# HG changeset patch # User haftmann # Date 1404497927 -7200 # Node ID cc97b347b301508a3b420c085d1c9a01f8268e7d # Parent de51a86fc9035bede32334f1806503687e6b1178 reduced name variants for assoc and commute on plus and mult diff -r de51a86fc903 -r cc97b347b301 NEWS --- a/NEWS Fri Jul 04 20:07:08 2014 +0200 +++ b/NEWS Fri Jul 04 20:18:47 2014 +0200 @@ -221,6 +221,22 @@ * The symbol "\" may be used within char or string literals to represent (Char Nibble0 NibbleA), i.e. ASCII newline. +* Reduced name variants for rules on associativity and commutativity: + + add_assoc ~> add.assoc + add_commute ~> add.commute + add_left_commute ~> add.left_commute + mult_assoc ~> mult.assoc + mult_commute ~> mult.commute + mult_left_commute ~> mult.left_commute + nat_add_assoc ~> add.assoc + nat_add_commute ~> add.commute + nat_add_left_commute ~> add.left_commute + nat_mult_assoc ~> mult.assoc + nat_mult_commute ~> mult.commute + eq_assoc ~> iff_assoc + eq_left_commute ~> iff_left_commute + * Qualified String.implode and String.explode. INCOMPATIBILITY. * Simplifier: Enhanced solver of preconditions of rewrite rules can diff -r de51a86fc903 -r cc97b347b301 src/Doc/Codegen/Further.thy --- a/src/Doc/Codegen/Further.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/Doc/Codegen/Further.thy Fri Jul 04 20:18:47 2014 +0200 @@ -205,7 +205,7 @@ interpretation %quote fun_power: power "\n (f :: 'a \ 'a). f ^^ n" where "power.powers (\n f. f ^^ n) = funpows" by unfold_locales - (simp_all add: fun_eq_iff funpow_mult mult_commute funpows_def) + (simp_all add: fun_eq_iff funpow_mult mult.commute funpows_def) text {* \noindent This additional equation is trivially proved by the diff -r de51a86fc903 -r cc97b347b301 src/Doc/Tutorial/Advanced/simp2.thy --- a/src/Doc/Tutorial/Advanced/simp2.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/Doc/Tutorial/Advanced/simp2.thy Fri Jul 04 20:18:47 2014 +0200 @@ -130,7 +130,7 @@ Each occurrence of an unknown is of the form $\Var{f}~x@1~\dots~x@n$, where the $x@i$ are distinct bound variables. Thus all ordinary rewrite rules, where all unknowns are -of base type, for example @{thm add_assoc}, are acceptable: if an unknown is +of base type, for example @{thm add.assoc}, are acceptable: if an unknown is of base type, it cannot have any arguments. Additionally, the rule @{text"(\x. ?P x \ ?Q x) = ((\x. ?P x) \ (\x. ?Q x))"} is also acceptable, in both directions: all arguments of the unknowns @{text"?P"} and diff -r de51a86fc903 -r cc97b347b301 src/Doc/Tutorial/Rules/Basic.thy --- a/src/Doc/Tutorial/Rules/Basic.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/Doc/Tutorial/Rules/Basic.thy Fri Jul 04 20:18:47 2014 +0200 @@ -64,11 +64,11 @@ text {* the subst method -@{thm[display] mult_commute} -\rulename{mult_commute} +@{thm[display] mult.commute} +\rulename{mult.commute} this would fail: -apply (simp add: mult_commute) +apply (simp add: mult.commute) *} @@ -76,7 +76,7 @@ txt{* @{subgoals[display,indent=0,margin=65]} *} -apply (subst mult_commute) +apply (subst mult.commute) txt{* @{subgoals[display,indent=0,margin=65]} *} @@ -84,7 +84,7 @@ (*exercise involving THEN*) lemma "\P x y z; Suc x < y\ \ f z = x*y" -apply (rule mult_commute [THEN ssubst]) +apply (rule mult.commute [THEN ssubst]) oops diff -r de51a86fc903 -r cc97b347b301 src/Doc/Tutorial/Types/Numbers.thy --- a/src/Doc/Tutorial/Types/Numbers.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/Doc/Tutorial/Types/Numbers.thy Fri Jul 04 20:18:47 2014 +0200 @@ -31,14 +31,14 @@ @{thm[display] add_2_eq_Suc'[no_vars]} \rulename{add_2_eq_Suc'} -@{thm[display] add_assoc[no_vars]} -\rulename{add_assoc} +@{thm[display] add.assoc[no_vars]} +\rulename{add.assoc} -@{thm[display] add_commute[no_vars]} -\rulename{add_commute} +@{thm[display] add.commute[no_vars]} +\rulename{add.commute} -@{thm[display] add_left_commute[no_vars]} -\rulename{add_left_commute} +@{thm[display] add.left_commute[no_vars]} +\rulename{add.left_commute} these form add_ac; similarly there is mult_ac *} diff -r de51a86fc903 -r cc97b347b301 src/Doc/Tutorial/document/numerics.tex --- a/src/Doc/Tutorial/document/numerics.tex Fri Jul 04 20:07:08 2014 +0200 +++ b/src/Doc/Tutorial/document/numerics.tex Fri Jul 04 20:18:47 2014 +0200 @@ -20,7 +20,7 @@ infix symbols. Algebraic properties are organized using type classes around algebraic concepts such as rings and fields; a property such as the commutativity of addition is a single theorem -(\isa{add_commute}) that applies to all numeric types. +(\isa{add.commute}) that applies to all numeric types. \index{linear arithmetic}% Many theorems involving numeric types can be proved automatically by @@ -441,11 +441,11 @@ the two expressions identical. \begin{isabelle} a\ +\ b\ +\ c\ =\ a\ +\ (b\ +\ c) -\rulenamedx{add_assoc}\isanewline +\rulenamedx{add.assoc}\isanewline a\ +\ b\ =\ b\ +\ a% -\rulenamedx{add_commute}\isanewline +\rulenamedx{add.commute}\isanewline a\ +\ (b\ +\ c)\ =\ b\ +\ (a\ +\ c) -\rulename{add_left_commute} +\rulename{add.left_commute} \end{isabelle} The name \isa{add_ac}\index{*add_ac (theorems)} refers to the list of all three theorems; similarly diff -r de51a86fc903 -r cc97b347b301 src/Doc/Tutorial/document/rules.tex --- a/src/Doc/Tutorial/document/rules.tex Fri Jul 04 20:07:08 2014 +0200 +++ b/src/Doc/Tutorial/document/rules.tex Fri Jul 04 20:18:47 2014 +0200 @@ -751,11 +751,11 @@ Now we wish to apply a commutative law: \begin{isabelle} ?m\ *\ ?n\ =\ ?n\ *\ ?m% -\rulename{mult_commute} +\rulename{mult.commute} \end{isabelle} Isabelle rejects our first attempt: \begin{isabelle} -apply (simp add: mult_commute) +apply (simp add: mult.commute) \end{isabelle} The simplifier notices the danger of looping and refuses to apply the rule.% @@ -764,9 +764,9 @@ is already smaller than \isa{y\ *\ x}.} % -The \isa{subst} method applies \isa{mult_commute} exactly once. +The \isa{subst} method applies \isa{mult.commute} exactly once. \begin{isabelle} -\isacommand{apply}\ (subst\ mult_commute)\isanewline +\isacommand{apply}\ (subst\ mult.commute)\isanewline \ 1.\ \isasymlbrakk P\ x\ y\ z;\ Suc\ x\ <\ y\isasymrbrakk \ \isasymLongrightarrow \ f\ z\ =\ y\ *\ x% \end{isabelle} @@ -775,7 +775,7 @@ \medskip This use of the \methdx{subst} method has the same effect as the command \begin{isabelle} -\isacommand{apply}\ (rule\ mult_commute [THEN ssubst]) +\isacommand{apply}\ (rule\ mult.commute [THEN ssubst]) \end{isabelle} The attribute \isa{THEN}, which combines two rules, is described in {\S}\ref{sec:THEN} below. The \methdx{subst} method is more powerful than @@ -1986,9 +1986,9 @@ % exercise worth including? E.g. find a difference between the two ways % of substituting. %\begin{exercise} -%In {\S}\ref{sec:subst} the method \isa{subst\ mult_commute} was applied. How +%In {\S}\ref{sec:subst} the method \isa{subst\ mult.commute} was applied. How %can we achieve the same effect using \isa{THEN} with the rule \isa{ssubst}? -%% answer rule (mult_commute [THEN ssubst]) +%% answer rule (mult.commute [THEN ssubst]) %\end{exercise} \subsection{Modifying a Theorem using {\tt\slshape OF}} diff -r de51a86fc903 -r cc97b347b301 src/HOL/Algebra/Coset.thy --- a/src/HOL/Algebra/Coset.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Algebra/Coset.thy Fri Jul 04 20:18:47 2014 +0200 @@ -815,7 +815,7 @@ "\finite(carrier G); subgroup H G\ \ card(rcosets H) * card(H) = order(G)" apply (simp (no_asm_simp) add: order_def rcosets_part_G [symmetric]) -apply (subst mult_commute) +apply (subst mult.commute) apply (rule card_partition) apply (simp add: rcosets_subset_PowG [THEN finite_subset]) apply (simp add: rcosets_part_G) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Algebra/Exponent.thy --- a/src/HOL/Algebra/Exponent.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Algebra/Exponent.thy Fri Jul 04 20:18:47 2014 +0200 @@ -80,7 +80,7 @@ lemma div_combine: fixes p::nat shows "[| prime p; ~ (p ^ (Suc r) dvd n); p^(a+r) dvd n*k |] ==> p ^ a dvd k" -by (metis add_Suc nat_add_commute prime_power_dvd_cases) +by (metis add_Suc add.commute prime_power_dvd_cases) (*Lemma for power_dvd_bound*) lemma Suc_le_power: "Suc 0 < p ==> Suc n <= p^n" @@ -185,7 +185,7 @@ text{*Main Combinatorial Argument*} lemma gcd_mult': fixes a::nat shows "gcd b (a * b) = b" -by (simp add: mult_commute[of a b]) +by (simp add: mult.commute[of a b]) lemma le_extend_mult: "[| c > 0; a <= b |] ==> a <= b * (c::nat)" apply (rule_tac P = "%x. x <= b * c" in subst) @@ -201,7 +201,7 @@ apply (drule less_imp_le [of a]) apply (drule le_imp_power_dvd) apply (drule_tac b = "p ^ r" in dvd_trans, assumption) -apply (metis diff_is_0_eq dvd_diffD1 gcd_dvd2_nat gcd_mult' gr0I le_extend_mult less_diff_conv nat_dvd_not_less nat_mult_commute not_add_less2 xt1(10)) +apply (metis diff_is_0_eq dvd_diffD1 gcd_dvd2_nat gcd_mult' gr0I le_extend_mult less_diff_conv nat_dvd_not_less mult.commute not_add_less2 xt1(10)) done lemma p_fac_forw: "[| (m::nat) > 0; k>0; k < p^a; (p^r) dvd (p^a)* m - k |] diff -r de51a86fc903 -r cc97b347b301 src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Algebra/Group.thy Fri Jul 04 20:18:47 2014 +0200 @@ -211,7 +211,7 @@ lemma (in monoid) nat_pow_pow: "x \ carrier G ==> (x (^) n) (^) m = x (^) (n * m::nat)" - by (induct m) (simp, simp add: nat_pow_mult add_commute) + by (induct m) (simp, simp add: nat_pow_mult add.commute) (* Jacobson defines submonoid here. *) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Algebra/IntRing.thy --- a/src/HOL/Algebra/IntRing.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Algebra/IntRing.thy Fri Jul 04 20:18:47 2014 +0200 @@ -271,11 +271,11 @@ then have "p dvd a \ p dvd b" by (metis prime prime_dvd_mult_eq_int) then show "(\x. a = x * int p) \ (\x. b = x * int p)" - by (metis dvd_def mult_commute) + by (metis dvd_def mult.commute) next assume "UNIV = {uu. EX x. uu = x * int p}" then obtain x where "1 = x * int p" by best - then have "\int p * x\ = 1" by (simp add: mult_commute) + then have "\int p * x\ = 1" by (simp add: mult.commute) then show False by (metis abs_of_nat int_1 of_nat_eq_iff abs_zmult_eq_1 one_not_prime_nat prime) qed diff -r de51a86fc903 -r cc97b347b301 src/HOL/Algebra/UnivPoly.thy --- a/src/HOL/Algebra/UnivPoly.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Algebra/UnivPoly.thy Fri Jul 04 20:18:47 2014 +0200 @@ -1353,7 +1353,7 @@ case 0 from R show ?case by simp next case Suc with R show ?case - by (simp del: monom_mult add: monom_mult [THEN sym] add_commute) + by (simp del: monom_mult add: monom_mult [THEN sym] add.commute) qed lemma (in ring_hom_cring) hom_pow [simp]: diff -r de51a86fc903 -r cc97b347b301 src/HOL/Code_Numeral.thy --- a/src/HOL/Code_Numeral.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Code_Numeral.thy Fri Jul 04 20:18:47 2014 +0200 @@ -494,7 +494,7 @@ show "k = integer_of_int (int_of_integer k)" by simp qed moreover have "2 * (j div 2) = j - j mod 2" - by (simp add: zmult_div_cancel mult_commute) + by (simp add: zmult_div_cancel mult.commute) ultimately show ?thesis by (auto simp add: split_def Let_def mod_integer_def nat_of_integer_def not_le nat_add_distrib [symmetric] Suc_nat_eq_nat_zadd1) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Complex.thy --- a/src/HOL/Complex.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Complex.thy Fri Jul 04 20:18:47 2014 +0200 @@ -118,7 +118,7 @@ show "scaleR (a + b) x = scaleR a x + scaleR b x" by (simp add: complex_eq_iff distrib_right) show "scaleR a (scaleR b x) = scaleR (a * b) x" - by (simp add: complex_eq_iff mult_assoc) + by (simp add: complex_eq_iff mult.assoc) show "scaleR 1 x = x" by (simp add: complex_eq_iff) show "scaleR a x * y = scaleR a (x * y)" @@ -190,7 +190,7 @@ by (simp add: divide_complex_def) lemma complex_i_mult_minus [simp]: "ii * (ii * x) = - x" - by (simp add: mult_assoc [symmetric]) + by (simp add: mult.assoc [symmetric]) lemma complex_i_not_zero [simp]: "ii \ 0" by (simp add: complex_eq_iff) @@ -298,14 +298,14 @@ apply (rule abs_sqrt_wlog [where x="Re z"]) apply (rule abs_sqrt_wlog [where x="Im z"]) apply (rule power2_le_imp_le) - apply (simp_all add: power2_sum add_commute sum_squares_bound real_sqrt_mult [symmetric]) + apply (simp_all add: power2_sum add.commute sum_squares_bound real_sqrt_mult [symmetric]) done text {* Properties of complex signum. *} lemma sgn_eq: "sgn z = z / complex_of_real (cmod z)" - by (simp add: sgn_div_norm divide_inverse scaleR_conv_of_real mult_commute) + by (simp add: sgn_div_norm divide_inverse scaleR_conv_of_real mult.commute) lemma Re_sgn [simp]: "Re(sgn z) = Re(z)/cmod z" by (simp add: complex_sgn_def divide_inverse) @@ -703,7 +703,7 @@ lemma complex_expi_Ex: "\a r. z = complex_of_real r * expi a" apply (insert rcis_Ex [of z]) -apply (auto simp add: expi_def rcis_def mult_assoc [symmetric]) +apply (auto simp add: expi_def rcis_def mult.assoc [symmetric]) apply (rule_tac x = "ii * complex_of_real a" in exI, auto) done diff -r de51a86fc903 -r cc97b347b301 src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Decision_Procs/Approximation.thy Fri Jul 04 20:18:47 2014 +0200 @@ -247,7 +247,7 @@ unfolding real_sqrt_mult[of _ "2 powr (?E mod 2)"] real_sqrt_abs2 by auto also have "\ < 2 powr (?E div 2) * 2 powr 1" by (rule mult_strict_left_mono, auto intro: E_mod_pow) - also have "\ = 2 powr (?E div 2 + 1)" unfolding add_commute[of _ 1] powr_add[symmetric] + also have "\ = 2 powr (?E div 2 + 1)" unfolding add.commute[of _ 1] powr_add[symmetric] by simp finally show ?thesis by auto qed @@ -386,8 +386,8 @@ { have "(x * lb_arctan_horner prec n 1 (x*x)) \ ?S n" using bounds(1) `0 \ real x` - unfolding power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric] - unfolding mult_commute[where 'a=real] mult_commute[of _ "2::nat"] power_mult power2_eq_square[of "real x"] + unfolding power_add power_one_right mult.assoc[symmetric] setsum_left_distrib[symmetric] + unfolding mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult power2_eq_square[of "real x"] by (auto intro!: mult_left_mono) also have "\ \ arctan x" using arctan_bounds .. finally have "(x * lb_arctan_horner prec n 1 (x*x)) \ arctan x" . } @@ -395,8 +395,8 @@ { have "arctan x \ ?S (Suc n)" using arctan_bounds .. also have "\ \ (x * ub_arctan_horner prec (Suc n) 1 (x*x))" using bounds(2)[of "Suc n"] `0 \ real x` - unfolding power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric] - unfolding mult_commute[where 'a=real] mult_commute[of _ "2::nat"] power_mult power2_eq_square[of "real x"] + unfolding power_add power_one_right mult.assoc[symmetric] setsum_left_distrib[symmetric] + unfolding mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult power2_eq_square[of "real x"] by (auto intro!: mult_left_mono) finally have "arctan x \ (x * ub_arctan_horner prec (Suc n) 1 (x*x))" . } ultimately show ?thesis by auto @@ -842,8 +842,8 @@ from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, OF `0 \ real (x * x)` f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps] show "?lb" and "?ub" using `0 \ real x` - unfolding power_add power_one_right mult_assoc[symmetric] setsum_left_distrib[symmetric] - unfolding mult_commute[where 'a=real] + unfolding power_add power_one_right mult.assoc[symmetric] setsum_left_distrib[symmetric] + unfolding mult.commute[where 'a=real] by (auto intro!: mult_left_mono simp add: power_mult power2_eq_square[of "real x"]) qed @@ -1597,7 +1597,7 @@ { fix n have "?a (Suc n) \ ?a n" unfolding inverse_eq_divide[symmetric] proof (rule mult_mono) show "0 \ x ^ Suc (Suc n)" by (auto simp add: `0 \ x`) - have "x ^ Suc (Suc n) \ x ^ Suc n * 1" unfolding power_Suc2 mult_assoc[symmetric] + have "x ^ Suc (Suc n) \ x ^ Suc n * 1" unfolding power_Suc2 mult.assoc[symmetric] by (rule mult_left_mono, fact less_imp_le[OF `x < 1`], auto simp: `0 \ x`) thus "x ^ Suc (Suc n) \ x ^ Suc n" by auto qed auto } @@ -1615,7 +1615,7 @@ let "?s n" = "-1^n * (1 / real (1 + n)) * (real x)^(Suc n)" - have "?lb \ setsum ?s {0 ..< 2 * ev}" unfolding power_Suc2 mult_assoc[symmetric] times_float.rep_eq setsum_left_distrib[symmetric] unfolding mult_commute[of "real x"] ev + have "?lb \ setsum ?s {0 ..< 2 * ev}" unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq setsum_left_distrib[symmetric] unfolding mult.commute[of "real x"] ev using horner_bounds(1)[where G="\ i k. Suc k" and F="\x. x" and f="\x. x" and lb="\n i k x. lb_ln_horner prec n k x" and ub="\n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*ev", OF `0 \ real x` refl lb_ln_horner.simps ub_ln_horner.simps] `0 \ real x` by (rule mult_right_mono) @@ -1623,7 +1623,7 @@ finally show "?lb \ ?ln" . have "?ln \ setsum ?s {0 ..< 2 * od + 1}" using ln_bounds(2)[OF `0 \ real x` `real x < 1`] by auto - also have "\ \ ?ub" unfolding power_Suc2 mult_assoc[symmetric] times_float.rep_eq setsum_left_distrib[symmetric] unfolding mult_commute[of "real x"] od + also have "\ \ ?ub" unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq setsum_left_distrib[symmetric] unfolding mult.commute[of "real x"] od using horner_bounds(2)[where G="\ i k. Suc k" and F="\x. x" and f="\x. x" and lb="\n i k x. lb_ln_horner prec n k x" and ub="\n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*od+1", OF `0 \ real x` refl lb_ln_horner.simps ub_ln_horner.simps] `0 \ real x` by (rule mult_right_mono) @@ -2961,7 +2961,7 @@ inverse (real (\ j \ {k.. {l .. u}" . } thus ?thesis using DERIV by blast diff -r de51a86fc903 -r cc97b347b301 src/HOL/Decision_Procs/Ferrack.thy --- a/src/HOL/Decision_Procs/Ferrack.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Decision_Procs/Ferrack.thy Fri Jul 04 20:18:47 2014 +0200 @@ -1407,7 +1407,7 @@ then obtain s m where smU: "(s,m) \ set (uset p)" and mx: "real m * x \ ?N a s" by blast from uset_l[OF lp] smU have mp: "real m > 0" by auto from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" - by (auto simp add: mult_commute) + by (auto simp add: mult.commute) thus ?thesis using smU by auto qed @@ -1423,7 +1423,7 @@ then obtain s m where smU: "(s,m) \ set (uset p)" and mx: "real m * x \ ?N a s" by blast from uset_l[OF lp] smU have mp: "real m > 0" by auto from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" - by (auto simp add: mult_commute) + by (auto simp add: mult.commute) thus ?thesis using smU by auto qed @@ -1674,7 +1674,7 @@ with uset_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" by auto let ?st = "Add (Mul m t) (Mul n s)" - from np mp have mnp: "real (2*n*m) > 0" by (simp add: mult_commute) + from np mp have mnp: "real (2*n*m) > 0" by (simp add: mult.commute) from tnb snb have st_nb: "numbound0 ?st" by simp have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" using mnp mp np by (simp add: algebra_simps add_divide_distrib) @@ -1690,7 +1690,7 @@ and px:"?I x (usubst p (Add (Mul l t) (Mul k s), 2*k*l))" with uset_l[OF lp] have tnb: "numbound0 t" and np:"real k > 0" and snb: "numbound0 s" and mp:"real l > 0" by auto let ?st = "Add (Mul l t) (Mul k s)" - from np mp have mnp: "real (2*k*l) > 0" by (simp add: mult_commute) + from np mp have mnp: "real (2*k*l) > 0" by (simp add: mult.commute) from tnb snb have st_nb: "numbound0 ?st" by simp from usubst_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto} ultimately show "?E" by blast @@ -1729,7 +1729,7 @@ (set U \ set U)"using mnz nnz th apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def) by (rule_tac x="(s,m)" in bexI,simp_all) - (rule_tac x="(t,n)" in bexI,simp_all add: mult_commute) + (rule_tac x="(t,n)" in bexI,simp_all add: mult.commute) next fix t n s m assume tnU: "(t,n) \ set U" and smU:"(s,m) \ set U" @@ -1778,7 +1778,7 @@ and snb: "numbound0 s" and mp:"m > 0" by auto let ?st= "Add (Mul m t) (Mul n s)" from np mp have mnp: "real (2*n*m) > 0" - by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) + by (simp add: mult.commute real_of_int_mult[symmetric] del: real_of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" using mp np by (simp add: algebra_simps add_divide_distrib) @@ -1806,7 +1806,7 @@ and snb: "numbound0 s" and mp:"m > 0" by auto let ?st= "Add (Mul m t) (Mul n s)" from np mp have mnp: "real (2*n*m) > 0" - by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) + by (simp add: mult.commute real_of_int_mult[symmetric] del: real_of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" using mp np by (simp add: algebra_simps add_divide_distrib) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Decision_Procs/MIR.thy Fri Jul 04 20:18:47 2014 +0200 @@ -4514,7 +4514,7 @@ then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real m * x \ ?N a s" by blast from \_l[OF lp] smU have mp: "real m > 0" by auto from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" - by (auto simp add: mult_commute) + by (auto simp add: mult.commute) thus ?thesis using smU by auto qed @@ -4530,7 +4530,7 @@ then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real m * x \ ?N a s" by blast from \_l[OF lp] smU have mp: "real m > 0" by auto from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" - by (auto simp add: mult_commute) + by (auto simp add: mult.commute) thus ?thesis using smU by auto qed @@ -4747,7 +4747,7 @@ with \_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" by auto let ?st = "Add (Mul m t) (Mul n s)" - from np mp have mnp: "real (2*n*m) > 0" by (simp add: mult_commute) + from np mp have mnp: "real (2*n*m) > 0" by (simp add: mult.commute) from tnb snb have st_nb: "numbound0 ?st" by simp have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" using mnp mp np by (simp add: algebra_simps add_divide_distrib) @@ -4763,7 +4763,7 @@ and px:"?I x (\ p (Add (Mul l t) (Mul k s), 2*k*l))" with \_l[OF lp] have tnb: "numbound0 t" and np:"real k > 0" and snb: "numbound0 s" and mp:"real l > 0" by auto let ?st = "Add (Mul l t) (Mul k s)" - from np mp have mnp: "real (2*k*l) > 0" by (simp add: mult_commute) + from np mp have mnp: "real (2*k*l) > 0" by (simp add: mult.commute) from tnb snb have st_nb: "numbound0 ?st" by simp from \_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto} ultimately show "?E" by blast @@ -4883,7 +4883,7 @@ from aU bU \_l[OF lrq] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" by (auto simp add: split_def) let ?st = "Add (Mul m t) (Mul n s)" from tnb snb have stnb: "numbound0 ?st" by simp - from np mp have mnp: "real (2*n*m) > 0" by (simp add: mult_commute) + from np mp have mnp: "real (2*n*m) > 0" by (simp add: mult.commute) from conjunct1[OF \_I[OF lrq mnp stnb, where bs="bs" and x="x"], symmetric] rqx have "\ x. ?I x ?rq" by auto thus "?E" @@ -4915,7 +4915,7 @@ (set U \ set U)"using mnz nnz th apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def) by (rule_tac x="(s,m)" in bexI,simp_all) - (rule_tac x="(t,n)" in bexI,simp_all add: mult_commute) + (rule_tac x="(t,n)" in bexI,simp_all add: mult.commute) next fix t n s m assume tnU: "(t,n) \ set U" and smU:"(s,m) \ set U" @@ -4964,7 +4964,7 @@ and snb: "numbound0 s" and mp:"m > 0" by auto let ?st= "Add (Mul m t) (Mul n s)" from np mp have mnp: "real (2*n*m) > 0" - by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) + by (simp add: mult.commute real_of_int_mult[symmetric] del: real_of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" using mp np by (simp add: algebra_simps add_divide_distrib) @@ -4992,7 +4992,7 @@ and snb: "numbound0 s" and mp:"m > 0" by auto let ?st= "Add (Mul m t) (Mul n s)" from np mp have mnp: "real (2*n*m) > 0" - by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) + by (simp add: mult.commute real_of_int_mult[symmetric] del: real_of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" using mp np by (simp add: algebra_simps add_divide_distrib) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy --- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Fri Jul 04 20:18:47 2014 +0200 @@ -428,7 +428,7 @@ apply (simp add: Let_def split_def field_simps) apply (simp add: Let_def split_def field_simps) apply (simp add: Let_def split_def field_simps) - apply (simp add: Let_def split_def mult_assoc distrib_left[symmetric]) + apply (simp add: Let_def split_def mult.assoc distrib_left[symmetric]) apply (simp add: Let_def split_def field_simps) apply (simp add: Let_def split_def field_simps) done @@ -1686,7 +1686,7 @@ assume xz: "x < -?e / ?c" then have "?c * x < - ?e" using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then have "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Eq (CNP 0 c e)))" @@ -1701,7 +1701,7 @@ assume xz: "x < -?e / ?c" then have "?c * x > - ?e" using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then have "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Eq (CNP 0 c e)))" @@ -1734,7 +1734,7 @@ assume xz: "x < -?e / ?c" then have "?c * x < - ?e" using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then have "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))" @@ -1749,7 +1749,7 @@ assume xz: "x < -?e / ?c" then have "?c * x > - ?e" using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then have "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))" @@ -1783,7 +1783,7 @@ assume xz: "x < -?e / ?c" then have "?c * x < - ?e" using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then have "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))" using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto @@ -1797,7 +1797,7 @@ assume xz: "x < -?e / ?c" then have "?c * x > - ?e" using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then have "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))" @@ -1830,7 +1830,7 @@ assume xz: "x < -?e / ?c" then have "?c * x < - ?e" using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then have "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))" @@ -1846,7 +1846,7 @@ assume xz: "x < -?e / ?c" then have "?c * x > - ?e" using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then have "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))" @@ -1899,7 +1899,7 @@ assume xz: "x > -?e / ?c" then have "?c * x > - ?e" using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then have "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))" @@ -1914,7 +1914,7 @@ assume xz: "x > -?e / ?c" then have "?c * x < - ?e" using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then have "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))" using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto @@ -1944,7 +1944,7 @@ assume xz: "x > -?e / ?c" then have "?c * x > - ?e" using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then have "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))" @@ -1959,7 +1959,7 @@ assume xz: "x > -?e / ?c" then have "?c * x < - ?e" using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then have "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))" @@ -1992,7 +1992,7 @@ assume xz: "x > -?e / ?c" then have "?c * x > - ?e" using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then have "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))" @@ -2007,7 +2007,7 @@ assume xz: "x > -?e / ?c" then have "?c * x < - ?e" using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then have "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))" @@ -2040,7 +2040,7 @@ assume xz: "x > -?e / ?c" then have "?c * x > - ?e" using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then have "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))" @@ -2055,7 +2055,7 @@ assume xz: "x > -?e / ?c" then have "?c * x < - ?e" using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] - by (simp add: mult_commute) + by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then have "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))" @@ -2144,7 +2144,7 @@ (Ipoly vs c > 0 \ Ipoly vs c * x \ - Itm vs (x#bs) s)" by blast then have "x \ (- Itm vs (x#bs) s) / Ipoly vs c" using divide_le_eq[of "- Itm vs (x#bs) s" "Ipoly vs c" x] - by (auto simp add: mult_commute) + by (auto simp add: mult.commute) then show ?thesis using csU by auto qed @@ -2188,7 +2188,7 @@ by blast then have "x \ (- Itm vs (x#bs) s) / Ipoly vs c" using le_divide_eq[of x "- Itm vs (x#bs) s" "Ipoly vs c"] - by (auto simp add: mult_commute) + by (auto simp add: mult.commute) then show ?thesis using csU by auto qed @@ -3811,7 +3811,7 @@ from msubst2[OF lp nn nn'(1), of x bs ] H(3) nn' have "\c\\<^sub>p\<^bsup>vs\<^esup> \ 0 \ \d\\<^sub>p\<^bsup>vs\<^esup> \ 0 \ Ifm vs ((- Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \d\\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p" - by (simp add: add_divide_distrib diff_divide_distrib mult_minus2_left mult_commute) + by (simp add: add_divide_distrib diff_divide_distrib mult_minus2_left mult.commute) } moreover { @@ -3828,7 +3828,7 @@ using H(3,4) by (simp_all add: polymul_norm n2) from msubst2[OF lp nn, of x bs ] H(3,4,5) have "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))" - by (simp add: diff_divide_distrib add_divide_distrib mult_minus2_left mult_commute) + by (simp add: diff_divide_distrib add_divide_distrib mult_minus2_left mult.commute) } ultimately show ?thesis by blast qed diff -r de51a86fc903 -r cc97b347b301 src/HOL/Decision_Procs/Polynomial_List.thy --- a/src/HOL/Decision_Procs/Polynomial_List.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Decision_Procs/Polynomial_List.thy Fri Jul 04 20:18:47 2014 +0200 @@ -121,7 +121,7 @@ thus ?case by auto next case (Cons b bs a) - thus ?case by (cases a) (simp_all add: add_commute) + thus ?case by (cases a) (simp_all add: add.commute) qed lemma (in comm_semiring_0) padd_assoc: "\b c. (a +++ b) +++ c = a +++ (b +++ c)" @@ -192,13 +192,13 @@ by simp lemma (in comm_semiring_0) poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x" - by (simp add: poly_mult mult_assoc) + by (simp add: poly_mult mult.assoc) lemma (in semiring_0) poly_mult_Nil2[simp]: "poly (p *** []) x = 0" by (induct p) auto lemma (in comm_semiring_1) poly_exp_add: "poly (p %^ (n + d)) x = poly( p %^ n *** p %^ d) x" - by (induct n) (auto simp add: poly_mult mult_assoc) + by (induct n) (auto simp add: poly_mult mult.assoc) subsection{*Key Property: if @{term "f(a) = 0"} then @{term "(x - a)"} divides @{term "p(x)"} *} @@ -445,7 +445,7 @@ lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a,1] \ poly []" apply (simp add: fun_eq) apply (rule_tac x = "minus one a" in exI) - apply (simp add: add_commute [of a]) + apply (simp add: add.commute [of a]) done lemma (in idom) poly_exp_prime_eq_zero: "poly ([a, 1] %^ n) \ poly []" @@ -517,7 +517,7 @@ lemma (in comm_semiring_1) poly_divides_trans: "p divides q \ q divides r \ p divides r" apply (simp add: divides_def, safe) apply (rule_tac x = "pmult qa qaa" in exI) - apply (auto simp add: poly_mult fun_eq mult_assoc) + apply (auto simp add: poly_mult fun_eq mult.assoc) done lemma (in comm_semiring_1) poly_divides_exp: "m \ n \ (p %^ m) divides (p %^ n)" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Decision_Procs/Rat_Pair.thy --- a/src/HOL/Decision_Procs/Rat_Pair.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Decision_Procs/Rat_Pair.thy Fri Jul 04 20:18:47 2014 +0200 @@ -464,7 +464,7 @@ qed lemma Nmul_commute: "isnormNum x \ isnormNum y \ x *\<^sub>N y = y *\<^sub>N x" - by (simp add: Nmul_def split_def Let_def gcd_commute_int mult_commute) + by (simp add: Nmul_def split_def Let_def gcd_commute_int mult.commute) lemma Nmul_assoc: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Deriv.thy Fri Jul 04 20:18:47 2014 +0200 @@ -558,7 +558,7 @@ lemma has_derivative_imp_has_field_derivative: "(f has_derivative D) F \ (\x. x * D' = D x) \ (f has_field_derivative D') F" unfolding has_field_derivative_def - by (rule has_derivative_eq_rhs[of f D]) (simp_all add: fun_eq_iff mult_commute) + by (rule has_derivative_eq_rhs[of f D]) (simp_all add: fun_eq_iff mult.commute) lemma has_field_derivative_imp_has_derivative: "(f has_field_derivative D) F \ (f has_derivative op * D) F" by (simp add: has_field_derivative_def) @@ -587,7 +587,7 @@ then obtain f' where *: "(f has_derivative f') (at x within s)" unfolding differentiable_def by auto then obtain c where "f' = (op * c)" - by (metis real_bounded_linear has_derivative_bounded_linear mult_commute fun_eq_iff) + by (metis real_bounded_linear has_derivative_bounded_linear mult.commute fun_eq_iff) with * show "\D. (f has_real_derivative D) (at x within s)" unfolding has_field_derivative_def by auto qed (auto simp: differentiable_def has_field_derivative_def) @@ -610,7 +610,7 @@ done lemma mult_commute_abs: "(\x. x * c) = op * (c::'a::ab_semigroup_mult)" - by (simp add: fun_eq_iff mult_commute) + by (simp add: fun_eq_iff mult.commute) subsection {* Derivatives *} @@ -732,7 +732,7 @@ "(f has_field_derivative d) (at x within s) \ (g has_field_derivative e) (at x within s)\ g x \ 0 \ ((\y. f y / g y) has_field_derivative (d * g x - (e * f x)) / (g x ^ Suc (Suc 0))) (at x within s)" - by (drule (2) DERIV_divide) (simp add: mult_commute) + by (drule (2) DERIV_divide) (simp add: mult.commute) lemma DERIV_power_Suc: "(f has_field_derivative D) (at x within s) \ @@ -765,7 +765,7 @@ lemma DERIV_chain: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ (f o g has_field_derivative Da * Db) (at x within s)" - by (drule (1) DERIV_chain', simp add: o_def mult_commute) + by (drule (1) DERIV_chain', simp add: o_def mult.commute) lemma DERIV_image_chain: "(f has_field_derivative Da) (at (g x) within (g ` s)) \ (g has_field_derivative Db) (at x within s) \ @@ -779,7 +779,7 @@ and "DERIV f x :> f'" and "f x \ s" shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" - by (metis (full_types) DERIV_chain' mult_commute assms) + by (metis (full_types) DERIV_chain' mult.commute assms) lemma DERIV_chain3: (*HAS_COMPLEX_DERIVATIVE_CHAIN_UNIV*) assumes "(\x. DERIV g x :> g'(x))" @@ -800,7 +800,7 @@ apply (drule_tac k="- a" in LIM_offset) apply simp apply (drule_tac k="a" in LIM_offset) -apply (simp add: add_commute) +apply (simp add: add.commute) done lemma DERIV_iff2: "(DERIV f x :> D) \ (\z. (f z - f x) / (z - x)) --x --> D" @@ -1215,7 +1215,7 @@ fixes f :: "real => real" shows "[|a \ b; \x. DERIV f x :> k |] ==> (f(b) - f(a))/(b-a) = k" apply (rule_tac c1 = "b-a" in mult_right_cancel [THEN iffD1]) -apply (auto dest!: DERIV_const_ratio_const simp add: mult_assoc) +apply (auto dest!: DERIV_const_ratio_const simp add: mult.assoc) done lemma real_average_minus_first [simp]: "((a + b) /2 - a) = (b-a)/(2::real)" @@ -1248,7 +1248,7 @@ hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) - ultimately show ?thesis using neq by (force simp add: add_commute) + ultimately show ?thesis using neq by (force simp add: add.commute) qed (* A function with positive derivative is increasing. diff -r de51a86fc903 -r cc97b347b301 src/HOL/Divides.thy --- a/src/HOL/Divides.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Divides.thy Fri Jul 04 20:18:47 2014 +0200 @@ -29,7 +29,7 @@ text {* @{const div} and @{const mod} *} lemma mod_div_equality2: "b * (a div b) + a mod b = a" - unfolding mult_commute [of b] + unfolding mult.commute [of b] by (rule mod_div_equality) lemma mod_div_equality': "a mod b + a div b * b = a" @@ -51,7 +51,7 @@ lemma div_mult_self2 [simp]: assumes "b \ 0" shows "(a + b * c) div b = c + a div b" - using assms div_mult_self1 [of b a c] by (simp add: mult_commute) + using assms div_mult_self1 [of b a c] by (simp add: mult.commute) lemma div_mult_self3 [simp]: assumes "b \ 0" @@ -74,7 +74,7 @@ "\ = (c + a div b) * b + (a + c * b) mod b" by (simp add: algebra_simps) finally have "a = a div b * b + (a + c * b) mod b" - by (simp add: add_commute [of a] add_assoc distrib_right) + by (simp add: add.commute [of a] add.assoc distrib_right) then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b" by (simp add: mod_div_equality) then show ?thesis by simp @@ -82,7 +82,7 @@ lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b" - by (simp add: mult_commute [of b]) + by (simp add: mult.commute [of b]) lemma mod_mult_self3 [simp]: "(c * b + a) mod b = a mod b" @@ -123,16 +123,16 @@ lemma div_add_self1 [simp]: assumes "b \ 0" shows "(b + a) div b = a div b + 1" - using assms div_mult_self1 [of b a 1] by (simp add: add_commute) + using assms div_mult_self1 [of b a 1] by (simp add: add.commute) lemma div_add_self2 [simp]: assumes "b \ 0" shows "(a + b) div b = a div b + 1" - using assms div_add_self1 [of b a] by (simp add: add_commute) + using assms div_add_self1 [of b a] by (simp add: add.commute) lemma mod_add_self1 [simp]: "(b + a) mod b = a mod b" - using mod_mult_self1 [of a 1 b] by (simp add: add_commute) + using mod_mult_self1 [of a 1 b] by (simp add: add.commute) lemma mod_add_self2 [simp]: "(a + b) mod b = a mod b" @@ -153,7 +153,7 @@ proof assume "b mod a = 0" with mod_div_equality [of b a] have "b div a * a = b" by simp - then have "b = a * (b div a)" unfolding mult_commute .. + then have "b = a * (b div a)" unfolding mult.commute .. then have "\c. b = a * c" .. then show "a dvd b" unfolding dvd_def . next @@ -161,7 +161,7 @@ then have "\c. b = a * c" unfolding dvd_def . then obtain c where "b = a * c" .. then have "b mod a = a * c mod a" by simp - then have "b mod a = c * a mod a" by (simp add: mult_commute) + then have "b mod a = c * a mod a" by (simp add: mult.commute) then show "b mod a = 0" by simp qed @@ -197,12 +197,12 @@ by (subst (2) mod_div_equality [of b a, symmetric]) (simp add:dvd_imp_mod_0) lemma dvd_mult_div_cancel: "a dvd b \ a * (b div a) = b" -by (drule dvd_div_mult_self) (simp add: mult_commute) +by (drule dvd_div_mult_self) (simp add: mult.commute) lemma dvd_div_mult: "a dvd b \ (b div a) * c = b * c div a" apply (cases "a = 0") apply simp -apply (auto simp: dvd_def mult_assoc) +apply (auto simp: dvd_def mult.assoc) done lemma div_dvd_div[simp]: @@ -211,8 +211,8 @@ apply simp apply (unfold dvd_def) apply auto - apply(blast intro:mult_assoc[symmetric]) -apply(fastforce simp add: mult_assoc) + apply(blast intro:mult.assoc[symmetric]) +apply(fastforce simp add: mult.assoc) done lemma dvd_mod_imp_dvd: "[| k dvd m mod n; k dvd n |] ==> k dvd m" @@ -332,7 +332,7 @@ apply (cases "y = 0", simp) apply (cases "z = 0", simp) apply (auto elim!: dvdE simp add: algebra_simps) - apply (subst mult_assoc [symmetric]) + apply (subst mult.assoc [symmetric]) apply (simp add: no_zero_divisors) done @@ -342,12 +342,12 @@ proof - from assms have "b div c * (a div 1) = b * a div (c * 1)" by (simp only: div_mult_div_if_dvd one_dvd) - then show ?thesis by (simp add: mult_commute) + then show ?thesis by (simp add: mult.commute) qed lemma div_mult_mult2 [simp]: "c \ 0 \ (a * c) div (b * c) = a div b" - by (drule div_mult_mult1) (simp add: mult_commute) + by (drule div_mult_mult1) (simp add: mult.commute) lemma div_mult_mult1_if [simp]: "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" @@ -368,7 +368,7 @@ lemma mod_mult_mult2: "(a * c) mod (b * c) = (a mod b) * c" - using mod_mult_mult1 [of c a b] by (simp add: mult_commute) + using mod_mult_mult1 [of c a b] by (simp add: mult.commute) lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)" by (fact mod_mult_mult2 [symmetric]) @@ -405,7 +405,7 @@ lemma dvd_div_div_eq_mult: assumes "a \ 0" "c \ 0" and "a dvd b" "c dvd d" shows "b div a = d div c \ b * c = a * d" - using assms by (auto simp add: mult_commute [of _ a] dvd_div_mult_self dvd_div_eq_mult div_mult_swap intro: sym) + using assms by (auto simp add: mult.commute [of _ a] dvd_div_mult_self dvd_div_eq_mult div_mult_swap intro: sym) end @@ -1065,7 +1065,7 @@ by (force simp add: divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN div_nat_unique]) lemma mod_mult2_eq: "a mod (b * c) = b * (a div b mod c) + a mod (b::nat)" -by (auto simp add: mult_commute divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN mod_nat_unique]) +by (auto simp add: mult.commute divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN mod_nat_unique]) subsubsection {* Further Facts about Quotient and Remainder *} @@ -1692,7 +1692,7 @@ fix a b :: int show "a div b * b + a mod b = a" using divmod_int_rel_div_mod [of a b] - unfolding divmod_int_rel_def by (simp add: mult_commute) + unfolding divmod_int_rel_def by (simp add: mult.commute) next fix a b c :: int assume "b \ 0" @@ -2022,7 +2022,7 @@ apply (subgoal_tac "b*q = r' - r + b'*q'") prefer 2 apply simp apply (simp (no_asm_simp) add: distrib_left) -apply (subst add_commute, rule add_less_le_mono, arith) +apply (subst add.commute, rule add_less_le_mono, arith) apply (rule mult_right_mono, auto) done @@ -2314,7 +2314,7 @@ "numeral (Num.Bit1 v) div numeral (Num.Bit0 w) = (numeral v div (numeral w :: int))" unfolding numeral.simps - unfolding mult_2 [symmetric] add_commute [of _ 1] + unfolding mult_2 [symmetric] add.commute [of _ 1] by (rule pos_zdiv_mult_2, simp) lemma pos_zmod_mult_2: @@ -2342,7 +2342,7 @@ "numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) = 2 * (numeral v mod numeral w) + (1::int)" unfolding numeral_Bit1 [of v] numeral_Bit0 [of w] - unfolding mult_2 [symmetric] add_commute [of _ 1] + unfolding mult_2 [symmetric] add.commute [of _ 1] by (rule pos_zmod_mult_2, simp) lemma zdiv_eq_0_iff: diff -r de51a86fc903 -r cc97b347b301 src/HOL/Fact.thy --- a/src/HOL/Fact.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Fact.thy Fri Jul 04 20:18:47 2014 +0200 @@ -230,7 +230,7 @@ apply auto apply (induct k rule: int_ge_induct) apply auto - apply (subst add_assoc [symmetric]) + apply (subst add.assoc [symmetric]) apply (subst fact_plus_one_int) apply auto apply (erule order_trans) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Fields.thy --- a/src/HOL/Fields.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Fields.thy Fri Jul 04 20:18:47 2014 +0200 @@ -51,7 +51,7 @@ assume ab: "a * b = 0" hence "0 = inverse a * (a * b) * inverse b" by simp also have "\ = (inverse a * a) * (b * inverse b)" - by (simp only: mult_assoc) + by (simp only: mult.assoc) also have "\ = 1" using a b by simp finally show False by simp qed @@ -81,7 +81,7 @@ proof - have "a \ 0" using ab by (cases "a = 0") simp_all moreover have "inverse a * (a * b) = inverse a" by (simp add: ab) - ultimately show ?thesis by (simp add: mult_assoc [symmetric]) + ultimately show ?thesis by (simp add: mult.assoc [symmetric]) qed lemma nonzero_inverse_minus_eq: @@ -110,7 +110,7 @@ shows "inverse (a * b) = inverse b * inverse a" proof - have "a * (b * inverse b) * inverse a = 1" using assms by simp - hence "a * b * (inverse b * inverse a) = 1" by (simp only: mult_assoc) + hence "a * b * (inverse b * inverse a) = 1" by (simp only: mult.assoc) thus ?thesis by (rule inverse_unique) qed @@ -126,7 +126,7 @@ proof assume neq: "b \ 0" { - hence "a = (a / b) * b" by (simp add: divide_inverse mult_assoc) + hence "a = (a / b) * b" by (simp add: divide_inverse mult.assoc) also assume "a / b = 1" finally show "a = b" by simp next @@ -154,7 +154,7 @@ by (simp add: divide_inverse) lemma times_divide_eq_right [simp]: "a * (b / c) = (a * b) / c" - by (simp add: divide_inverse mult_assoc) + by (simp add: divide_inverse mult.assoc) lemma minus_divide_left: "- (a / b) = (-a) / b" by (simp add: divide_inverse) @@ -175,7 +175,7 @@ proof - assume [simp]: "c \ 0" have "a = b / c \ a * c = (b / c) * c" by simp - also have "... \ a * c = b" by (simp add: divide_inverse mult_assoc) + also have "... \ a * c = b" by (simp add: divide_inverse mult.assoc) finally show ?thesis . qed @@ -183,7 +183,7 @@ proof - assume [simp]: "c \ 0" have "b / c = a \ (b / c) * c = a * c" by simp - also have "... \ b = a * c" by (simp add: divide_inverse mult_assoc) + also have "... \ b = a * c" by (simp add: divide_inverse mult.assoc) finally show ?thesis . qed @@ -194,10 +194,10 @@ using nonzero_neg_divide_eq_eq[of b a c] by auto lemma divide_eq_imp: "c \ 0 \ b = a * c \ b / c = a" - by (simp add: divide_inverse mult_assoc) + by (simp add: divide_inverse mult.assoc) lemma eq_divide_imp: "c \ 0 \ a * c = b \ a = b / c" - by (drule sym) (simp add: divide_inverse mult_assoc) + by (drule sym) (simp add: divide_inverse mult.assoc) lemma add_divide_eq_iff [field_simps]: "z \ 0 \ x + y / z = (x * z + y) / z" @@ -298,7 +298,7 @@ fix a :: 'a assume "a \ 0" thus "inverse a * a = 1" by (rule field_inverse) - thus "a * inverse a = 1" by (simp only: mult_commute) + thus "a * inverse a = 1" by (simp only: mult.commute) next fix a b :: 'a show "a / b = a * inverse b" by (rule field_divide_inverse) @@ -325,7 +325,7 @@ lemma nonzero_mult_divide_mult_cancel_right [simp]: "\b \ 0; c \ 0\ \ (a * c) / (b * c) = a / b" -by (simp add: mult_commute [of _ c]) +by (simp add: mult.commute [of _ c]) lemma times_divide_eq_left [simp]: "(b / c) * a = (b * a) / c" by (simp add: divide_inverse mult_ac) @@ -346,7 +346,7 @@ also have "\ = (x * z + y * w) / (y * z)" by (simp only: add_divide_distrib) finally show ?thesis - by (simp only: mult_commute) + by (simp only: mult.commute) qed text{*Special Cancellation Simprules for Division*} @@ -406,7 +406,7 @@ lemma inverse_divide [simp]: "inverse (a / b) = b / a" - by (simp add: divide_inverse mult_commute) + by (simp add: divide_inverse mult.commute) text {* Calculations with fractions *} @@ -433,7 +433,7 @@ lemma divide_divide_eq_left [simp]: "(a / b) / c = a / (b * c)" - by (simp add: divide_inverse mult_assoc) + by (simp add: divide_inverse mult.assoc) lemma divide_divide_times_eq: "(x / y) / (z / w) = (x * w) / (y * z)" @@ -538,7 +538,7 @@ by (simp add: apos invle less_imp_le mult_left_mono) hence "(a * inverse a) * b \ (a * inverse b) * b" by (simp add: bpos less_imp_le mult_right_mono) - thus "b \ a" by (simp add: mult_assoc apos bpos less_imp_not_eq2) + thus "b \ a" by (simp add: mult.assoc apos bpos less_imp_not_eq2) qed lemma inverse_positive_imp_positive: @@ -669,7 +669,7 @@ hence "(a \ b/c) = (a*c \ (b/c)*c)" by (simp add: mult_le_cancel_right less_not_sym [OF less] del: times_divide_eq) also have "... = (a*c \ b)" - by (simp add: less_imp_not_eq2 [OF less] divide_inverse mult_assoc) + by (simp add: less_imp_not_eq2 [OF less] divide_inverse mult.assoc) finally show ?thesis . qed @@ -679,7 +679,7 @@ hence "(a \ b/c) = ((b/c)*c \ a*c)" by (simp add: mult_le_cancel_right less_not_sym [OF less] del: times_divide_eq) also have "... = (b \ a*c)" - by (simp add: less_imp_not_eq [OF less] divide_inverse mult_assoc) + by (simp add: less_imp_not_eq [OF less] divide_inverse mult.assoc) finally show ?thesis . qed @@ -689,7 +689,7 @@ hence "(a < b/c) = (a*c < (b/c)*c)" by (simp add: mult_less_cancel_right_disj less_not_sym [OF less] del: times_divide_eq) also have "... = (a*c < b)" - by (simp add: less_imp_not_eq2 [OF less] divide_inverse mult_assoc) + by (simp add: less_imp_not_eq2 [OF less] divide_inverse mult.assoc) finally show ?thesis . qed @@ -699,7 +699,7 @@ hence "(a < b/c) = ((b/c)*c < a*c)" by (simp add: mult_less_cancel_right_disj less_not_sym [OF less] del: times_divide_eq) also have "... = (b < a*c)" - by (simp add: less_imp_not_eq [OF less] divide_inverse mult_assoc) + by (simp add: less_imp_not_eq [OF less] divide_inverse mult.assoc) finally show ?thesis . qed @@ -709,7 +709,7 @@ hence "(b/c < a) = ((b/c)*c < a*c)" by (simp add: mult_less_cancel_right_disj less_not_sym [OF less] del: times_divide_eq) also have "... = (b < a*c)" - by (simp add: less_imp_not_eq2 [OF less] divide_inverse mult_assoc) + by (simp add: less_imp_not_eq2 [OF less] divide_inverse mult.assoc) finally show ?thesis . qed @@ -719,7 +719,7 @@ hence "(b/c < a) = (a*c < (b/c)*c)" by (simp add: mult_less_cancel_right_disj less_not_sym [OF less] del: times_divide_eq) also have "... = (a*c < b)" - by (simp add: less_imp_not_eq [OF less] divide_inverse mult_assoc) + by (simp add: less_imp_not_eq [OF less] divide_inverse mult.assoc) finally show ?thesis . qed @@ -729,7 +729,7 @@ hence "(b/c \ a) = ((b/c)*c \ a*c)" by (simp add: mult_le_cancel_right less_not_sym [OF less] del: times_divide_eq) also have "... = (b \ a*c)" - by (simp add: less_imp_not_eq2 [OF less] divide_inverse mult_assoc) + by (simp add: less_imp_not_eq2 [OF less] divide_inverse mult.assoc) finally show ?thesis . qed @@ -739,7 +739,7 @@ hence "(b/c \ a) = (a*c \ (b/c)*c)" by (simp add: mult_le_cancel_right less_not_sym [OF less] del: times_divide_eq) also have "... = (a*c \ b)" - by (simp add: less_imp_not_eq [OF less] divide_inverse mult_assoc) + by (simp add: less_imp_not_eq [OF less] divide_inverse mult.assoc) finally show ?thesis . qed @@ -1049,7 +1049,7 @@ lemma divide_left_mono_neg: "a <= b ==> c <= 0 ==> 0 < a * b ==> c / a <= c / b" apply (drule divide_left_mono [of _ _ "- c"]) - apply (auto simp add: mult_commute) + apply (auto simp add: mult.commute) done lemma inverse_le_iff: "inverse a \ inverse b \ (0 < a * b \ b \ a) \ (a * b \ 0 \ a \ b)" diff -r de51a86fc903 -r cc97b347b301 src/HOL/GCD.thy --- a/src/HOL/GCD.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/GCD.thy Fri Jul 04 20:18:47 2014 +0200 @@ -410,7 +410,7 @@ apply (rule_tac n = k in coprime_dvd_mult_nat) apply (simp add: gcd_assoc_nat) apply (simp add: gcd_commute_nat) - apply (simp_all add: mult_commute) + apply (simp_all add: mult.commute) done lemma gcd_mult_cancel_int: @@ -432,9 +432,9 @@ with `coprime a d` have "a dvd b" by (simp add: coprime_dvd_mult_iff_nat) from `?lhs` have "b dvd a * c" by (auto intro: dvdI dest: sym) with `coprime b c` have "b dvd a" by (simp add: coprime_dvd_mult_iff_nat) - from `?lhs` have "c dvd d * b" by (auto intro: dvdI dest: sym simp add: mult_commute) + from `?lhs` have "c dvd d * b" by (auto intro: dvdI dest: sym simp add: mult.commute) with `coprime b c` have "c dvd d" by (simp add: coprime_dvd_mult_iff_nat gcd_commute_nat) - from `?lhs` have "d dvd c * a" by (auto intro: dvdI dest: sym simp add: mult_commute) + from `?lhs` have "d dvd c * a" by (auto intro: dvdI dest: sym simp add: mult.commute) with `coprime a d` have "d dvd c" by (simp add: coprime_dvd_mult_iff_nat gcd_commute_nat) from `a dvd b` `b dvd a` have "a = b" by (rule Nat.dvd.antisym) moreover from `c dvd d` `d dvd c` have "c = d" by (rule Nat.dvd.antisym) @@ -456,7 +456,7 @@ lemma gcd_add2_nat [simp]: "gcd (m::nat) (m + n) = gcd m n" apply (subst (1 2) gcd_commute_nat) - apply (subst add_commute) + apply (subst add.commute) apply simp done @@ -496,16 +496,16 @@ done lemma gcd_add1_int [simp]: "gcd ((m::int) + n) n = gcd m n" -by (metis gcd_red_int mod_add_self1 add_commute) +by (metis gcd_red_int mod_add_self1 add.commute) lemma gcd_add2_int [simp]: "gcd m ((m::int) + n) = gcd m n" -by (metis gcd_add1_int gcd_commute_int add_commute) +by (metis gcd_add1_int gcd_commute_int add.commute) lemma gcd_add_mult_nat: "gcd (m::nat) (k * m + n) = gcd m n" by (metis mod_mult_self3 gcd_commute_nat gcd_red_nat) lemma gcd_add_mult_int: "gcd (m::int) (k * m + n) = gcd m n" -by (metis gcd_commute_int gcd_red_int mod_mult_self1 add_commute) +by (metis gcd_commute_int gcd_red_int mod_mult_self1 add.commute) (* to do: differences, and all variations of addition rules @@ -778,7 +778,7 @@ by (auto simp:div_gcd_coprime_nat) hence "gcd ((a div gcd a b)^n * (gcd a b)^n) ((b div gcd a b)^n * (gcd a b)^n) = (gcd a b)^n" - apply (subst (1 2) mult_commute) + apply (subst (1 2) mult.commute) apply (subst gcd_mult_distrib_nat [symmetric]) apply simp done @@ -820,10 +820,10 @@ from ab'(1) have "a' dvd a" unfolding dvd_def by blast with dc have th0: "a' dvd b*c" using dvd_trans[of a' a "b*c"] by simp from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto - hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc) + hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult.assoc) with z have th_1: "a' dvd b' * c" by auto from coprime_dvd_mult_nat[OF ab'(3)] th_1 - have thc: "a' dvd c" by (subst (asm) mult_commute, blast) + have thc: "a' dvd c" by (subst (asm) mult.commute, blast) from ab' have "a = ?g*a'" by algebra with thb thc have ?thesis by blast } ultimately show ?thesis by blast @@ -844,10 +844,10 @@ with dc have th0: "a' dvd b*c" using dvd_trans[of a' a "b*c"] by simp from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto - hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc) + hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult.assoc) with z have th_1: "a' dvd b' * c" by auto from coprime_dvd_mult_int[OF ab'(3)] th_1 - have thc: "a' dvd c" by (subst (asm) mult_commute, blast) + have thc: "a' dvd c" by (subst (asm) mult.commute, blast) from ab' have "a = ?g*a'" by algebra with thb thc have ?thesis by blast } ultimately show ?thesis by blast @@ -869,13 +869,13 @@ from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n" by (simp add: ab'(1,2)[symmetric]) hence "?g^n*a'^n dvd ?g^n *b'^n" - by (simp only: power_mult_distrib mult_commute) + by (simp only: power_mult_distrib mult.commute) with zn z n have th0:"a'^n dvd b'^n" by auto have "a' dvd a'^n" by (simp add: m) with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp - hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute) + hence th1: "a' dvd b'^m * b'" by (simp add: m mult.commute) from coprime_dvd_mult_nat[OF coprime_exp_nat [OF ab'(3), of m]] th1 - have "a' dvd b'" by (subst (asm) mult_commute, blast) + have "a' dvd b'" by (subst (asm) mult.commute, blast) hence "a'*?g dvd b'*?g" by simp with ab'(1,2) have ?thesis by simp } ultimately show ?thesis by blast @@ -897,14 +897,14 @@ from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n" by (simp add: ab'(1,2)[symmetric]) hence "?g^n*a'^n dvd ?g^n *b'^n" - by (simp only: power_mult_distrib mult_commute) + by (simp only: power_mult_distrib mult.commute) with zn z n have th0:"a'^n dvd b'^n" by auto have "a' dvd a'^n" by (simp add: m) with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp - hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute) + hence th1: "a' dvd b'^m * b'" by (simp add: m mult.commute) from coprime_dvd_mult_int[OF coprime_exp_int [OF ab'(3), of m]] th1 - have "a' dvd b'" by (subst (asm) mult_commute, blast) + have "a' dvd b'" by (subst (asm) mult.commute, blast) hence "a'*?g dvd b'*?g" by simp with ab'(1,2) have ?thesis by simp } ultimately show ?thesis by blast @@ -922,7 +922,7 @@ proof- from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'" unfolding dvd_def by blast - from mr n' have "m dvd n'*n" by (simp add: mult_commute) + from mr n' have "m dvd n'*n" by (simp add: mult.commute) hence "m dvd n'" using coprime_dvd_mult_iff_nat[OF mn] by simp then obtain k where k: "n' = m*k" unfolding dvd_def by blast from n' k show ?thesis unfolding dvd_def by auto @@ -934,7 +934,7 @@ proof- from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'" unfolding dvd_def by blast - from mr n' have "m dvd n'*n" by (simp add: mult_commute) + from mr n' have "m dvd n'*n" by (simp add: mult.commute) hence "m dvd n'" using coprime_dvd_mult_iff_int[OF mn] by simp then obtain k where k: "n' = m*k" unfolding dvd_def by blast from n' k show ?thesis unfolding dvd_def by auto @@ -1218,14 +1218,14 @@ from H (3) have "d + (b - 1) * (b*x) = d + (b - 1) * (a*y + d)" by simp hence "d + (b - 1) * a * y + (b - 1) * d = d + (b - 1) * b * x" - by (simp only: mult_assoc distrib_left) + by (simp only: mult.assoc distrib_left) hence "a * ((b - 1) * y) + d * (b - 1 + 1) = d + x*b*(b - 1)" by algebra hence "a * ((b - 1) * y) = d + x*b*(b - 1) - d*b" using bp by simp hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)" by (simp only: diff_add_assoc[OF dble, of d, symmetric]) hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d" - by (simp only: diff_mult_distrib2 add_commute mult_ac) + by (simp only: diff_mult_distrib2 add.commute mult_ac) hence ?thesis using H(1,2) apply - apply (rule exI[where x=d], simp) @@ -1674,7 +1674,7 @@ apply(auto simp add:inj_on_def) apply (metis coprime_dvd_mult_iff_nat dvd.neq_le_trans dvd_triv_left) apply (metis gcd_semilattice_nat.inf_commute coprime_dvd_mult_iff_nat - dvd.neq_le_trans dvd_triv_right mult_commute) + dvd.neq_le_trans dvd_triv_right mult.commute) done text{* Nitpick: *} diff -r de51a86fc903 -r cc97b347b301 src/HOL/Groups.thy --- a/src/HOL/Groups.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Groups.thy Fri Jul 04 20:18:47 2014 +0200 @@ -158,6 +158,8 @@ end +hide_fact add_assoc + class ab_semigroup_add = semigroup_add + assumes add_commute [algebra_simps, field_simps]: "a + b = b + a" begin @@ -165,13 +167,15 @@ sublocale add!: abel_semigroup plus by default (fact add_commute) -lemmas add_left_commute [algebra_simps, field_simps] = add.left_commute +declare add.left_commute [algebra_simps, field_simps] -theorems add_ac = add_assoc add_commute add_left_commute +theorems add_ac = add.assoc add.commute add.left_commute end -theorems add_ac = add_assoc add_commute add_left_commute +hide_fact add_commute + +theorems add_ac = add.assoc add.commute add.left_commute class semigroup_mult = times + assumes mult_assoc [algebra_simps, field_simps]: "(a * b) * c = a * (b * c)" @@ -182,6 +186,8 @@ end +hide_fact mult_assoc + class ab_semigroup_mult = semigroup_mult + assumes mult_commute [algebra_simps, field_simps]: "a * b = b * a" begin @@ -189,13 +195,15 @@ sublocale mult!: abel_semigroup times by default (fact mult_commute) -lemmas mult_left_commute [algebra_simps, field_simps] = mult.left_commute +declare mult.left_commute [algebra_simps, field_simps] -theorems mult_ac = mult_assoc mult_commute mult_left_commute +theorems mult_ac = mult.assoc mult.commute mult.left_commute end -theorems mult_ac = mult_assoc mult_commute mult_left_commute +hide_fact mult_commute + +theorems mult_ac = mult.assoc mult.commute mult.left_commute class monoid_add = zero + semigroup_add + assumes add_0_left: "0 + a = a" @@ -325,7 +333,7 @@ next fix a b c :: 'a assume "b + a = c + a" - then have "a + b = a + c" by (simp only: add_commute) + then have "a + b = a + c" by (simp only: add.commute) then show "b = c" by (rule add_imp_eq) qed @@ -349,7 +357,7 @@ assumes "a + b = 0" shows "- a = b" proof - have "- a = - a + (a + b)" using assms by simp - also have "\ = b" by (simp add: add_assoc [symmetric]) + also have "\ = b" by (simp add: add.assoc [symmetric]) finally show ?thesis . qed @@ -381,36 +389,36 @@ fix a b c :: 'a assume "a + b = a + c" then have "- a + a + b = - a + a + c" - unfolding add_assoc by simp + unfolding add.assoc by simp then show "b = c" by simp next fix a b c :: 'a assume "b + a = c + a" then have "b + a + - a = c + a + - a" by simp - then show "b = c" unfolding add_assoc by simp + then show "b = c" unfolding add.assoc by simp qed lemma minus_add_cancel [simp]: "- a + (a + b) = b" - by (simp add: add_assoc [symmetric]) + by (simp add: add.assoc [symmetric]) lemma add_minus_cancel [simp]: "a + (- a + b) = b" - by (simp add: add_assoc [symmetric]) + by (simp add: add.assoc [symmetric]) lemma diff_add_cancel [simp]: "a - b + b = a" - by (simp only: diff_conv_add_uminus add_assoc) simp + by (simp only: diff_conv_add_uminus add.assoc) simp lemma add_diff_cancel [simp]: "a + b - b = a" - by (simp only: diff_conv_add_uminus add_assoc) simp + by (simp only: diff_conv_add_uminus add.assoc) simp lemma minus_add: "- (a + b) = - b + - a" proof - have "(a + b) + (- b + - a) = 0" - by (simp only: add_assoc add_minus_cancel) simp + by (simp only: add.assoc add_minus_cancel) simp then show "- (a + b) = - b + - a" by (rule minus_unique) qed @@ -419,7 +427,7 @@ "a - b = 0 \ a = b" proof assume "a - b = 0" - have "a = (a - b) + b" by (simp add: add_assoc) + have "a = (a - b) + b" by (simp add: add.assoc) also have "\ = b" using `a - b = 0` by simp finally show "a = b" . next @@ -484,7 +492,7 @@ next assume "a + b = 0" moreover have "a + (b + - b) = (a + b) + - b" - by (simp only: add_assoc) + by (simp only: add.assoc) ultimately show "a = - b" by simp qed @@ -502,15 +510,15 @@ lemma minus_diff_eq [simp]: "- (a - b) = b - a" - by (simp only: neg_eq_iff_add_eq_0 diff_conv_add_uminus add_assoc minus_add_cancel) simp + by (simp only: neg_eq_iff_add_eq_0 diff_conv_add_uminus add.assoc minus_add_cancel) simp lemma add_diff_eq [algebra_simps, field_simps]: "a + (b - c) = (a + b) - c" - by (simp only: diff_conv_add_uminus add_assoc) + by (simp only: diff_conv_add_uminus add.assoc) lemma diff_add_eq_diff_diff_swap: "a - (b + c) = a - c - b" - by (simp only: diff_conv_add_uminus add_assoc minus_add) + by (simp only: diff_conv_add_uminus add.assoc minus_add) lemma diff_eq_eq [algebra_simps, field_simps]: "a - b = c \ a = c + b" @@ -522,7 +530,7 @@ lemma diff_diff_eq2 [algebra_simps, field_simps]: "a - (b - c) = (a + c) - b" - by (simp only: diff_conv_add_uminus add_assoc) simp + by (simp only: diff_conv_add_uminus add.assoc) simp lemma diff_eq_diff_eq: "a - b = c - d \ a = b \ c = d" @@ -543,13 +551,13 @@ fix a b c :: 'a assume "a + b = a + c" then have "- a + a + b = - a + a + c" - by (simp only: add_assoc) + by (simp only: add.assoc) then show "b = c" by simp qed lemma uminus_add_conv_diff [simp]: "- a + b = b - a" - by (simp add: add_commute) + by (simp add: add.commute) lemma minus_add_distrib [simp]: "- (a + b) = - a + - b" @@ -595,13 +603,13 @@ lemma add_right_mono: "a \ b \ a + c \ b + c" -by (simp add: add_commute [of _ c] add_left_mono) +by (simp add: add.commute [of _ c] add_left_mono) text {* non-strict, in both arguments *} lemma add_mono: "a \ b \ c \ d \ a + c \ b + d" apply (erule add_right_mono [THEN order_trans]) - apply (simp add: add_commute add_left_mono) + apply (simp add: add.commute add_left_mono) done end @@ -616,7 +624,7 @@ lemma add_strict_right_mono: "a < b \ a + c < b + c" -by (simp add: add_commute [of _ c] add_strict_left_mono) +by (simp add: add.commute [of _ c] add_strict_left_mono) text{*Strict monotonicity in both arguments*} lemma add_strict_mono: @@ -665,7 +673,7 @@ lemma add_less_imp_less_right: "a + c < b + c \ a < b" apply (rule add_less_imp_less_left [of c]) -apply (simp add: add_commute) +apply (simp add: add.commute) done lemma add_less_cancel_left [simp]: @@ -682,7 +690,7 @@ lemma add_le_cancel_right [simp]: "a + c \ b + c \ a \ b" - by (simp add: add_commute [of a c] add_commute [of b c]) + by (simp add: add.commute [of a c] add.commute [of b c]) lemma add_le_imp_le_right: "a + c \ b + c \ a \ b" @@ -721,45 +729,45 @@ lemma add_diff_assoc: "c + (b - a) = c + b - a" - using `a \ b` by (auto simp add: le_iff_add add_left_commute [of c]) + using `a \ b` by (auto simp add: le_iff_add add.left_commute [of c]) lemma add_diff_assoc2: "b - a + c = b + c - a" - using `a \ b` by (auto simp add: le_iff_add add_assoc) + using `a \ b` by (auto simp add: le_iff_add add.assoc) lemma diff_add_assoc: "c + b - a = c + (b - a)" - using `a \ b` by (simp add: add_commute add_diff_assoc) + using `a \ b` by (simp add: add.commute add_diff_assoc) lemma diff_add_assoc2: "b + c - a = b - a + c" - using `a \ b`by (simp add: add_commute add_diff_assoc) + using `a \ b`by (simp add: add.commute add_diff_assoc) lemma diff_diff_right: "c - (b - a) = c + a - b" - by (simp add: add_diff_inverse add_diff_cancel_left [of a c "b - a", symmetric] add_commute) + by (simp add: add_diff_inverse add_diff_cancel_left [of a c "b - a", symmetric] add.commute) lemma diff_add: "b - a + a = b" - by (simp add: add_commute add_diff_inverse) + by (simp add: add.commute add_diff_inverse) lemma le_add_diff: "c \ b + c - a" - by (auto simp add: add_commute diff_add_assoc2 le_iff_add) + by (auto simp add: add.commute diff_add_assoc2 le_iff_add) lemma le_imp_diff_is_add: "a \ b \ b - a = c \ b = c + a" - by (auto simp add: add_commute add_diff_inverse) + by (auto simp add: add.commute add_diff_inverse) lemma le_diff_conv2: "c \ b - a \ c + a \ b" (is "?P \ ?Q") proof assume ?P then have "c + a \ b - a + a" by (rule add_right_mono) - then show ?Q by (simp add: add_diff_inverse add_commute) + then show ?Q by (simp add: add_diff_inverse add.commute) next assume ?Q - then have "a + c \ a + (b - a)" by (simp add: add_diff_inverse add_commute) + then have "a + c \ a + (b - a)" by (simp add: add_diff_inverse add.commute) then show ?P by simp qed @@ -860,7 +868,7 @@ lemma add_increasing2: "0 \ c \ b \ a \ b \ a + c" - by (simp add: add_increasing add_commute [of a]) + by (simp add: add_increasing add.commute [of a]) lemma add_strict_increasing: "0 < a \ b \ c \ b < a + c" @@ -883,7 +891,7 @@ fix a b c :: 'a assume "c + a \ c + b" hence "(-c) + (c + a) \ (-c) + (c + b)" by (rule add_left_mono) - hence "((-c) + c) + a \ ((-c) + c) + b" by (simp only: add_assoc) + hence "((-c) + c) + a \ ((-c) + c) + b" by (simp only: add.assoc) thus "a \ b" by simp qed diff -r de51a86fc903 -r cc97b347b301 src/HOL/Groups_Big.thy --- a/src/HOL/Groups_Big.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Groups_Big.thy Fri Jul 04 20:18:47 2014 +0200 @@ -958,7 +958,7 @@ shows "setsum (\i. card {j\T. R i j}) S = k * card T" (is "?l = ?r") proof- have "?l = setsum (\i. k) T" by (rule setsum_multicount_gen) (auto simp: assms) - also have "\ = ?r" by (simp add: mult_commute) + also have "\ = ?r" by (simp add: mult.commute) finally show ?thesis by auto qed diff -r de51a86fc903 -r cc97b347b301 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/HOL.thy Fri Jul 04 20:18:47 2014 +0200 @@ -977,10 +977,10 @@ by blast lemma eq_ac: - shows eq_commute: "(a=b) = (b=a)" - and eq_left_commute: "(P=(Q=R)) = (Q=(P=R))" - and eq_assoc: "((P=Q)=R) = (P=(Q=R))" by (iprover, blast+) -lemma neq_commute: "(a~=b) = (b~=a)" by iprover + shows eq_commute: "a = b \ b = a" + and iff_left_commute: "(P \ (Q \ R)) \ (Q \ (P \ R))" + and iff_assoc: "((P \ Q) \ R) \ (P \ (Q \ R))" by (iprover, blast+) +lemma neq_commute: "a \ b \ b \ a" by iprover lemma conj_comms: shows conj_commute: "(P&Q) = (Q&P)" diff -r de51a86fc903 -r cc97b347b301 src/HOL/HOLCF/FOCUS/Buffer_adm.thy --- a/src/HOL/HOLCF/FOCUS/Buffer_adm.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/HOLCF/FOCUS/Buffer_adm.thy Fri Jul 04 20:18:47 2014 +0200 @@ -123,7 +123,7 @@ apply (rule allI) apply (induct_tac "i") apply ( simp) -apply (simp add: add_commute) +apply (simp add: add.commute) apply (intro strip) apply (subst BufAC_Cmt_F_def3) apply (drule_tac P="%x. x" in BufAC_Cmt_F_def3 [THEN subst]) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Hahn_Banach/Function_Norm.thy --- a/src/HOL/Hahn_Banach/Function_Norm.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Hahn_Banach/Function_Norm.thy Fri Jul 04 20:18:47 2014 +0200 @@ -141,7 +141,7 @@ then show "0 \ inverse \x\" by (rule order_less_imp_le) qed also have "\ = c * (\x\ * inverse \x\)" - by (rule Groups.mult_assoc) + by (rule Groups.mult.assoc) also from gt have "\x\ \ 0" by simp then have "\x\ * inverse \x\ = 1" by simp diff -r de51a86fc903 -r cc97b347b301 src/HOL/Hahn_Banach/Vector_Space.thy --- a/src/HOL/Hahn_Banach/Vector_Space.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Hahn_Banach/Vector_Space.thy Fri Jul 04 20:18:47 2014 +0200 @@ -209,9 +209,9 @@ proof from y have "y = 0 + y" by simp also from x y have "\ = (- x + x) + y" by simp - also from x y have "\ = - x + (x + y)" by (simp add: add_assoc) + also from x y have "\ = - x + (x + y)" by (simp add: add.assoc) also assume "x + y = x + z" - also from x z have "- x + (x + z) = - x + x + z" by (simp add: add_assoc) + also from x z have "- x + (x + z) = - x + x + z" by (simp add: add.assoc) also from x z have "\ = z" by simp finally show "y = z" . next @@ -228,7 +228,7 @@ by (simp only: add_assoc [symmetric]) lemma mult_left_commute: "x \ V \ a \ b \ x = b \ a \ x" - by (simp add: mult_commute mult_assoc2) + by (simp add: mult.commute mult_assoc2) lemma mult_zero_uniq: assumes x: "x \ V" "x \ 0" and ax: "a \ x = 0" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Hoare_Parallel/RG_Examples.thy --- a/src/HOL/Hoare_Parallel/RG_Examples.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Hoare_Parallel/RG_Examples.thy Fri Jul 04 20:18:47 2014 +0200 @@ -52,7 +52,7 @@ apply force apply(erule le_less_trans2) apply(case_tac t,simp+) - apply (simp add:add_commute) + apply (simp add:add.commute) apply(simp add: add_le_mono) apply(rule Basic) apply simp @@ -61,7 +61,7 @@ apply simp apply(erule le_less_trans2) apply(case_tac t,simp+) - apply (simp add:add_commute) + apply (simp add:add.commute) apply(rule add_le_mono, auto) done diff -r de51a86fc903 -r cc97b347b301 src/HOL/IMP/ACom.thy --- a/src/HOL/IMP/ACom.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/IMP/ACom.thy Fri Jul 04 20:18:47 2014 +0200 @@ -104,7 +104,7 @@ apply(induction c arbitrary: f p) apply (auto simp: anno_def nth_append nth_Cons numeral_eq_Suc shift_def split: nat.split) - apply (metis add_Suc_right add_diff_inverse nat_add_commute) + apply (metis add_Suc_right add_diff_inverse add.commute) apply(rule_tac f=f in arg_cong) apply arith apply (metis less_Suc_eq) diff -r de51a86fc903 -r cc97b347b301 src/HOL/IMP/Abs_Int0.thy --- a/src/HOL/IMP/Abs_Int0.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/IMP/Abs_Int0.thy Fri Jul 04 20:18:47 2014 +0200 @@ -317,7 +317,7 @@ "m_s S X = (\ x \ X. m(S x))" lemma m_s_h: "finite X \ m_s S X \ h * card X" -by(simp add: m_s_def) (metis nat_mult_commute of_nat_id setsum_bounded[OF h]) +by(simp add: m_s_def) (metis mult.commute of_nat_id setsum_bounded[OF h]) fun m_o :: "'av st option \ vname set \ nat" ("m\<^sub>o") where "m_o (Some S) X = m_s S X" | diff -r de51a86fc903 -r cc97b347b301 src/HOL/IMP/Abs_Int1.thy --- a/src/HOL/IMP/Abs_Int1.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/IMP/Abs_Int1.thy Fri Jul 04 20:18:47 2014 +0200 @@ -109,7 +109,7 @@ "m_s S X = (\ x \ X. m(fun S x))" lemma m_s_h: "finite X \ m_s S X \ h * card X" -by(simp add: m_s_def) (metis nat_mult_commute of_nat_id setsum_bounded[OF h]) +by(simp add: m_s_def) (metis mult.commute of_nat_id setsum_bounded[OF h]) definition m_o :: "'av st option \ vname set \ nat" ("m\<^sub>o") where "m_o opt X = (case opt of None \ h * card X + 1 | Some S \ m_s S X)" diff -r de51a86fc903 -r cc97b347b301 src/HOL/IMP/Abs_Int_Den/Abs_Int_den1_ivl.thy --- a/src/HOL/IMP/Abs_Int_Den/Abs_Int_den1_ivl.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/IMP/Abs_Int_Den/Abs_Int_den1_ivl.thy Fri Jul 04 20:18:47 2014 +0200 @@ -193,7 +193,7 @@ proof case goal1 thus ?case by(auto simp add: filter_plus_ivl_def) - (metis rep_minus_ivl add_diff_cancel add_commute)+ + (metis rep_minus_ivl add_diff_cancel add.commute)+ next case goal2 thus ?case by(cases a1, cases a2, diff -r de51a86fc903 -r cc97b347b301 src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ivl_ITP.thy --- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ivl_ITP.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int2_ivl_ITP.thy Fri Jul 04 20:18:47 2014 +0200 @@ -208,7 +208,7 @@ next case goal2 thus ?case by(auto simp add: filter_plus_ivl_def) - (metis gamma_minus_ivl add_diff_cancel add_commute)+ + (metis gamma_minus_ivl add_diff_cancel add.commute)+ next case goal3 thus ?case by(cases a1, cases a2, diff -r de51a86fc903 -r cc97b347b301 src/HOL/Int.thy --- a/src/HOL/Int.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Int.thy Fri Jul 04 20:18:47 2014 +0200 @@ -567,7 +567,7 @@ "(1 + z + z < 0) = (z < (0::int))" proof (cases z) case (nonneg n) - thus ?thesis by (simp add: linorder_not_less add_assoc add_increasing + thus ?thesis by (simp add: linorder_not_less add.assoc add_increasing le_imp_0_less [THEN order_less_imp_le]) next case (neg n) @@ -585,7 +585,7 @@ case (nonneg n) have le: "0 \ z+z" by (simp add: nonneg add_increasing) thus ?thesis using le_imp_0_less [OF le] - by (auto simp add: add_assoc) + by (auto simp add: add.assoc) next case (neg n) show ?thesis @@ -594,7 +594,7 @@ have "(0::int) < 1 + (int n + int n)" by (simp add: le_imp_0_less add_increasing) also have "... = - (1 + z + z)" - by (simp add: neg add_assoc [symmetric]) + by (simp add: neg add.assoc [symmetric]) also have "... = 0" by (simp add: eq) finally have "0<0" .. thus False by blast @@ -1079,7 +1079,7 @@ lemma zmult_eq_1_iff: "(m*n = (1::int)) = ((m = 1 & n = 1) | (m = -1 & n = -1))" apply (rule iffI) apply (frule pos_zmult_eq_1_iff_lemma) - apply (simp add: mult_commute [of m]) + apply (simp add: mult.commute [of m]) apply (frule pos_zmult_eq_1_iff_lemma, auto) done @@ -1238,7 +1238,7 @@ lemma zdvd_antisym_nonneg: "0 <= m ==> 0 <= n ==> m dvd n ==> n dvd m ==> m = (n::int)" apply (simp add: dvd_def, auto) - apply (auto simp add: mult_assoc zero_le_mult_iff zmult_eq_1_iff) + apply (auto simp add: mult.assoc zero_le_mult_iff zmult_eq_1_iff) done lemma zdvd_antisym_abs: assumes "(a::int) dvd b" and "b dvd a" @@ -1251,7 +1251,7 @@ from `b dvd a` obtain k' where k':"a = b*k'" unfolding dvd_def by blast from k k' have "a = a*k*k'" by simp with mult_cancel_left1[where c="a" and b="k*k'"] - have kk':"k*k' = 1" using `a\0` by (simp add: mult_assoc) + have kk':"k*k' = 1" using `a\0` by (simp add: mult.assoc) hence "k = 1 \ k' = 1 \ k = -1 \ k' = -1" by (simp add: zmult_eq_1_iff) thus ?thesis using k k' by auto qed @@ -1301,7 +1301,7 @@ proof- from d obtain h where h: "k*n = k*m * h" unfolding dvd_def by blast {assume "n \ m*h" hence "k* n \ k* (m*h)" using kz by simp - with h have False by (simp add: mult_assoc)} + with h have False by (simp add: mult.assoc)} hence "n = m * h" by blast thus ?thesis by simp qed diff -r de51a86fc903 -r cc97b347b301 src/HOL/Isar_Examples/Fibonacci.thy --- a/src/HOL/Isar_Examples/Fibonacci.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Isar_Examples/Fibonacci.thy Fri Jul 04 20:18:47 2014 +0200 @@ -92,7 +92,7 @@ proof - assume "0 < n" then have "gcd (n * k + m) n = gcd n (m mod n)" - by (simp add: gcd_non_0_nat add_commute) + by (simp add: gcd_non_0_nat add.commute) also from `0 < n` have "\ = gcd m n" by (simp add: gcd_non_0_nat) finally show ?thesis . diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/BigO.thy --- a/src/HOL/Library/BigO.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/BigO.thy Fri Jul 04 20:18:47 2014 +0200 @@ -489,7 +489,7 @@ shows "c \ 0 \ f \ O(\x. c * f x)" apply (simp add: bigo_def) apply (rule_tac x = "abs (inverse c)" in exI) - apply (simp add: abs_mult [symmetric] mult_assoc [symmetric]) + apply (simp add: abs_mult [symmetric] mult.assoc [symmetric]) done lemma bigo_const_mult4: @@ -517,10 +517,10 @@ apply (simp add: func_times) apply (auto intro!: simp add: bigo_def elt_set_times_def func_times) apply (rule_tac x = "\y. inverse c * x y" in exI) - apply (simp add: mult_assoc [symmetric] abs_mult) + apply (simp add: mult.assoc [symmetric] abs_mult) apply (rule_tac x = "abs (inverse c) * ca" in exI) apply (rule allI) - apply (subst mult_assoc) + apply (subst mult.assoc) apply (rule mult_left_mono) apply (erule spec) apply force @@ -613,7 +613,7 @@ apply (rule_tac x = c in exI) apply (rule allI)+ apply (subst abs_mult)+ - apply (subst mult_left_commute) + apply (subst mult.left_commute) apply (rule mult_left_mono) apply (erule spec) apply (rule abs_ge_zero) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Code_Target_Int.thy --- a/src/HOL/Library/Code_Target_Int.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Code_Target_Int.thy Fri Jul 04 20:18:47 2014 +0200 @@ -103,7 +103,7 @@ from mod_div_equality have *: "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp show ?thesis by (simp add: Let_def divmod_int_mod_div of_int_add [symmetric]) - (simp add: * mult_commute) + (simp add: * mult.commute) qed declare of_int_code_if [code] diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Code_Target_Nat.thy --- a/src/HOL/Library/Code_Target_Nat.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Code_Target_Nat.thy Fri Jul 04 20:18:47 2014 +0200 @@ -112,7 +112,7 @@ from mod_div_equality have *: "of_nat n = of_nat (n div 2 * 2 + n mod 2)" by simp show ?thesis by (simp add: Let_def divmod_nat_div_mod of_nat_add [symmetric]) - (simp add: * mult_commute of_nat_mult add_commute) + (simp add: * mult.commute of_nat_mult add.commute) qed declare of_nat_code_if [code] diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Convex.thy --- a/src/HOL/Library/Convex.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Convex.thy Fri Jul 04 20:18:47 2014 +0200 @@ -210,7 +210,7 @@ have "(\j \ {1..2}. ?u j *\<^sub>R ?x j) = \ *\<^sub>R x + (1 - \) *\<^sub>R y" by auto then have "(1 - \) *\<^sub>R y + \ *\<^sub>R x \ s" - using s by (auto simp:add_commute) + using s by (auto simp:add.commute) } then show "convex s" unfolding convex_alt by auto @@ -525,7 +525,7 @@ using a_nonneg a1 asms by blast have "f (\ j \ insert i s. a j *\<^sub>R y j) = f ((\ j \ s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" using setsum.insert[of s i "\ j. a j *\<^sub>R y j", OF `finite s` `i \ s`] asms - by (auto simp only:add_commute) + by (auto simp only:add.commute) also have "\ = f (((1 - a i) * inverse (1 - a i)) *\<^sub>R (\ j \ s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" using i0 by auto also have "\ = f ((1 - a i) *\<^sub>R (\ j \ s. (a j * inverse (1 - a i)) *\<^sub>R y j) + a i *\<^sub>R y i)" @@ -535,7 +535,7 @@ by (auto simp: divide_inverse) also have "\ \ (1 - a i) *\<^sub>R f ((\ j \ s. ?a j *\<^sub>R y j)) + a i * f (y i)" using conv[of "y i" "(\ j \ s. ?a j *\<^sub>R y j)" "a i", OF yai(1) asum yai(2) ai1] - by (auto simp add:add_commute) + by (auto simp add:add.commute) also have "\ \ (1 - a i) * (\ j \ s. ?a j * f (y j)) + a i * f (y i)" using add_right_mono[OF mult_left_mono[of _ _ "1 - a i", OF asum_le less_imp_le[OF i0]], of "a i * f (y i)"] by simp @@ -768,7 +768,7 @@ by auto then have f''0: "\z::real. z > 0 \ DERIV (\ z. - 1 / (ln b * z)) z :> 1 / (ln b * z * z)" - unfolding inverse_eq_divide by (auto simp add: mult_assoc) + unfolding inverse_eq_divide by (auto simp add: mult.assoc) have f''_ge0: "\z::real. z > 0 \ 1 / (ln b * z * z) \ 0" using `b > 1` by (auto intro!:less_imp_le) from f''_ge0_imp_convex[OF pos_is_convex, diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Extended_Nat.thy --- a/src/HOL/Library/Extended_Nat.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Extended_Nat.thy Fri Jul 04 20:18:47 2014 +0200 @@ -175,7 +175,7 @@ by (simp_all add: eSuc_plus_1 add_ac) lemma iadd_Suc_right: "m + eSuc n = eSuc (m + n)" - by (simp only: add_commute[of m] iadd_Suc) + by (simp only: add.commute[of m] iadd_Suc) lemma iadd_is_0: "(m + n = (0::enat)) = (m = 0 \ n = 0)" by (cases m, cases n, simp_all add: zero_enat_def) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Extended_Real.thy --- a/src/HOL/Library/Extended_Real.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Extended_Real.thy Fri Jul 04 20:18:47 2014 +0200 @@ -701,7 +701,7 @@ lemma ereal_mult_strict_left_mono: "a < b \ 0 < c \ c < (\::ereal) \ c * a < c * b" using ereal_mult_strict_right_mono - by (simp add: mult_commute[of c]) + by (simp add: mult.commute[of c]) lemma ereal_mult_right_mono: fixes a b c :: ereal @@ -717,7 +717,7 @@ fixes a b c :: ereal shows "a \ b \ 0 \ c \ c * a \ c * b" using ereal_mult_right_mono - by (simp add: mult_commute[of c]) + by (simp add: mult.commute[of c]) lemma zero_less_one_ereal[simp]: "0 \ (1::ereal)" by (simp add: one_ereal_def zero_ereal_def) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Float.thy --- a/src/HOL/Library/Float.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Float.thy Fri Jul 04 20:18:47 2014 +0200 @@ -844,7 +844,7 @@ hence "1 \ real m * inverse ?S" using assms False powr_realpow[of 2 "nat (-e)"] by (auto simp: powr_minus) hence "1 * ?S \ real m * inverse ?S * ?S" by (rule mult_right_mono, auto) - hence "?S \ real m" unfolding mult_assoc by auto + hence "?S \ real m" unfolding mult.assoc by auto hence "?S \ m" unfolding real_of_int_le_iff[symmetric] by auto from this bitlen_bounds[OF `0 < m`, THEN conjunct2] have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF `1 < 2`, symmetric] by (rule order_le_less_trans) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Formal_Power_Series.thy --- a/src/HOL/Library/Formal_Power_Series.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Formal_Power_Series.thy Fri Jul 04 20:18:47 2014 +0200 @@ -128,14 +128,14 @@ proof fix a b c :: "'a fps" show "a + b + c = a + (b + c)" - by (simp add: fps_ext add_assoc) + by (simp add: fps_ext add.assoc) qed instance fps :: (ab_semigroup_add) ab_semigroup_add proof fix a b :: "'a fps" show "a + b = b + a" - by (simp add: fps_ext add_commute) + by (simp add: fps_ext add.commute) qed lemma fps_mult_assoc_lemma: @@ -143,7 +143,7 @@ and f :: "nat \ nat \ nat \ 'a::comm_monoid_add" shows "(\j=0..k. \i=0..j. f i (j - i) (n - j)) = (\j=0..k. \i=0..k - j. f j i (n - j - i))" - by (induct k) (simp_all add: Suc_diff_le setsum.distrib add_assoc) + by (induct k) (simp_all add: Suc_diff_le setsum.distrib add.assoc) instance fps :: (semiring_0) semigroup_mult proof @@ -155,7 +155,7 @@ (\j=0..n. \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_right_distrib setsum_left_distrib mult.assoc) qed qed @@ -174,7 +174,7 @@ have "(\i=0..n. a$i * b$(n - i)) = (\i=0..n. a$(n - i) * b$i)" by (rule fps_mult_commute_lemma) then show "(a * b) $ n = (b * a) $ n" - by (simp add: fps_mult_nth mult_commute) + by (simp add: fps_mult_nth mult.commute) qed qed @@ -397,7 +397,7 @@ lemma X_mult_right_nth[simp]: "((f :: 'a::comm_semiring_1 fps) * X) $n = (if n = 0 then 0 else f $ (n - 1))" - by (metis X_mult_nth mult_commute) + by (metis X_mult_nth mult.commute) lemma X_power_iff: "X^k = Abs_fps (\n. if n = k then 1::'a::comm_ring_1 else 0)" proof (induct k) @@ -419,14 +419,14 @@ "(X^k * (f :: 'a::comm_ring_1 fps)) $n = (if n < k then 0 else f $ (n - k))" apply (induct k arbitrary: n) apply simp - unfolding power_Suc mult_assoc + unfolding power_Suc mult.assoc apply (case_tac n) apply auto done lemma X_power_mult_right_nth: "((f :: 'a::comm_ring_1 fps) * X^k) $n = (if n < k then 0 else f $ (n - k))" - by (metis X_power_mult_nth mult_commute) + by (metis X_power_mult_nth mult.commute) subsection{* Formal Power series form a metric space *} @@ -666,7 +666,7 @@ shows "inverse f * f = 1" proof - have c: "inverse f * f = f * inverse f" - by (simp add: mult_commute) + by (simp add: mult.commute) from f0 have ifn: "\n. inverse f $ n = natfun_inverse f n" by (simp add: fps_inverse_def) from f0 have th0: "(inverse f * f) $ 0 = 1" @@ -809,7 +809,7 @@ setsum (\i. f $ (n + 1 - i) * g $ i) ?Zn1" by (rule setsum.reindex_bij_witness[where i="op - (n + 1)" and j="op - (n + 1)"]) auto have "(f * ?D g + ?D f * g)$n = (?D g * f + ?D f * g)$n" - by (simp only: mult_commute) + by (simp only: mult.commute) also have "\ = (\i = 0..n. ?g i)" by (simp add: fps_mult_nth setsum.distrib[symmetric]) also have "\ = setsum ?h {0..n+1}" @@ -948,7 +948,7 @@ lemma fps_nth_deriv_mult_const_right[simp]: "fps_nth_deriv n (f * fps_const (c::'a::comm_ring_1)) = fps_nth_deriv n f * fps_const c" - using fps_nth_deriv_linear[of n "c" f 0 0] by (simp add: mult_commute) + using fps_nth_deriv_linear[of n "c" f 0 0] by (simp add: mult.commute) lemma fps_nth_deriv_setsum: "fps_nth_deriv n (setsum f S) = setsum (\i. fps_nth_deriv n (f i :: 'a::comm_ring_1 fps)) S" @@ -1026,7 +1026,7 @@ { assume m0: "m \ 0" have "a ^k $ m = (a^l * a) $m" - by (simp add: k mult_commute) + by (simp add: k mult.commute) also have "\ = (\i = 0..m. a ^ l $ i * a $ (m - i))" by (simp add: fps_mult_nth) also have "\ = 0" @@ -1114,7 +1114,7 @@ unfolding power_mult_distrib[symmetric] apply (rule ssubst[where t = "a * inverse a" and s= 1]) apply simp_all - apply (subst mult_commute) + apply (subst mult.commute) apply (rule inverse_mult_eq_1[OF a0]) done } @@ -1144,7 +1144,7 @@ have "(inverse a)\<^sup>2 * fps_deriv a + fps_deriv (inverse a) = 0" unfolding power2_eq_square apply (simp add: field_simps) - apply (simp add: mult_assoc[symmetric]) + apply (simp add: mult.assoc[symmetric]) done then have "(inverse a)\<^sup>2 * fps_deriv a + fps_deriv (inverse a) - fps_deriv a * (inverse a)\<^sup>2 = 0 - fps_deriv a * (inverse a)\<^sup>2" @@ -1194,7 +1194,7 @@ lemma inverse_mult_eq_1': assumes f0: "f$0 \ (0::'a::field)" shows "f * inverse f= 1" - by (metis mult_commute inverse_mult_eq_1 f0) + by (metis mult.commute inverse_mult_eq_1 f0) lemma fps_divide_deriv: fixes a :: "'a::field fps" @@ -1398,7 +1398,7 @@ by simp have "a /?X = ?X * Abs_fps (\n::nat. setsum (op $ a) {0..n}) * inverse ?X" using fps_divide_X_minus1_setsum_lemma[of a, symmetric] th0 - by (simp add: fps_divide_def mult_assoc) + by (simp add: fps_divide_def mult.assoc) also have "\ = (inverse ?X * ?X) * Abs_fps (\n::nat. setsum (op $ a) {0..n}) " by (simp add: mult_ac) finally show ?thesis @@ -2289,14 +2289,14 @@ done also have "\ = setsum (\i. of_nat (i + 1) * a$(i+1) * (setsum (\j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" unfolding fps_deriv_nth - by (rule setsum.reindex_cong [of Suc]) (auto simp add: mult_assoc) + by (rule setsum.reindex_cong [of Suc]) (auto simp add: mult.assoc) finally have th0: "(fps_deriv (a oo b))$n = setsum (\i. of_nat (i + 1) * a$(i+1) * (setsum (\j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" . have "(((fps_deriv a) oo b) * (fps_deriv b))$n = setsum (\i. (fps_deriv b)$ (n - i) * ((fps_deriv a) oo b)$i) {0..n}" unfolding fps_mult_nth by (simp add: mult_ac) also have "\ = setsum (\i. setsum (\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_right_distrib mult.assoc apply (rule setsum.cong) apply (rule refl) apply (rule setsum.mono_neutral_left) @@ -2503,7 +2503,7 @@ apply (simp add: fps_mult_nth fps_compose_nth setsum_product) apply (rule setsum.cong) apply (rule refl) - apply (simp add: setsum.cartesian_product mult_assoc) + apply (simp add: setsum.cartesian_product mult.assoc) apply (rule setsum.mono_neutral_right[OF f]) apply (simp add: subset_eq) apply presburger @@ -2689,11 +2689,11 @@ 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_right_distrib mult.assoc) lemma fps_const_mult_apply_right: "(a oo b) * fps_const (c::'a::comm_semiring_1) = (fps_const c * a) oo b" - by (auto simp add: fps_const_mult_apply_left mult_commute) + by (auto simp add: fps_const_mult_apply_left mult.commute) lemma fps_compose_assoc: assumes c0: "c$0 = (0::'a::idom)" @@ -2704,11 +2704,11 @@ fix n have "?l$n = (setsum (\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_right_distrib mult.assoc fps_setsum_nth) also have "\ = ((setsum (\i. fps_const (a$i) * b^i) {0..n}) oo c)$n" by (simp add: fps_compose_setsum_distrib) also have "\ = ?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_left_distrib mult.assoc) apply (rule setsum.cong) apply (rule refl) apply (rule setsum.mono_neutral_right) @@ -3052,7 +3052,7 @@ using a by (simp add: fps_const_inverse eq fps_divide_def fps_inverse_mult) then have "fps_deriv ?l = fps_deriv ?r" - by (simp add: fps_deriv_L add_commute fps_divide_def divide_inverse) + by (simp add: fps_deriv_L add.commute fps_divide_def divide_inverse) then show ?thesis unfolding fps_deriv_eq_iff by (simp add: L_nth fps_inv_def) qed @@ -3094,7 +3094,7 @@ have x10: "?x1 $ 0 \ 0" by simp have "?l = ?r \ inverse ?x1 * ?l = inverse ?x1 * ?r" by simp also have "\ \ ?da = (fps_const c * a) / ?x1" - apply (simp only: fps_divide_def mult_assoc[symmetric] inverse_mult_eq_1[OF x10]) + apply (simp only: fps_divide_def mult.assoc[symmetric] inverse_mult_eq_1[OF x10]) apply (simp add: field_simps) done finally have eq: "?l = ?r \ ?lhs" by simp @@ -3119,7 +3119,7 @@ case (Suc m) then show ?case unfolding th0 apply (simp add: field_simps del: of_nat_Suc) - unfolding mult_assoc[symmetric] gbinomial_mult_1 + unfolding mult.assoc[symmetric] gbinomial_mult_1 apply (simp add: field_simps) done qed @@ -3135,12 +3135,12 @@ { assume h: ?rhs have th00: "\x y. x * (a$0 * y) = a$0 * (x*y)" - by (simp add: mult_commute) + by (simp add: mult.commute) have "?l = ?r" apply (subst h) apply (subst (2) h) apply (clarsimp simp add: fps_eq_iff field_simps) - unfolding mult_assoc[symmetric] th00 gbinomial_mult_1 + unfolding mult.assoc[symmetric] th00 gbinomial_mult_1 apply (simp add: field_simps gbinomial_mult_1) done } @@ -3181,7 +3181,7 @@ have th: "?r$0 \ 0" by simp have th': "fps_deriv (inverse ?r) = fps_const (- 1) * inverse ?r / (1 + X)" by (simp add: fps_inverse_deriv[OF th] fps_divide_def - power2_eq_square mult_commute fps_const_neg[symmetric] del: fps_const_neg) + power2_eq_square mult.commute fps_const_neg[symmetric] del: fps_const_neg) have eq: "inverse ?r $ 0 = 1" by (simp add: fps_inverse_def) from iffD1[OF fps_binomial_ODE_unique[of "inverse (1 + X)" "- 1"] th'] eq diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Fraction_Field.thy --- a/src/HOL/Library/Fraction_Field.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Fraction_Field.thy Fri Jul 04 20:18:47 2014 +0200 @@ -252,7 +252,7 @@ assume "q \ 0" then show "inverse q * q = 1" by (cases q rule: Fract_cases_nonzero) - (simp_all add: fract_expand eq_fract mult_commute) + (simp_all add: fract_expand eq_fract mult.commute) next fix q r :: "'a fract" show "q / r = q * inverse r" by (simp add: divide_fract_def) @@ -398,7 +398,7 @@ by (simp only: less_fract_def) show "q \ r \ r \ q" by (induct q, induct r) - (simp add: mult_commute, rule linorder_linear) + (simp add: mult.commute, rule linorder_linear) qed end diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Fundamental_Theorem_Algebra.thy --- a/src/HOL/Library/Fundamental_Theorem_Algebra.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Fundamental_Theorem_Algebra.thy Fri Jul 04 20:18:47 2014 +0200 @@ -1006,9 +1006,9 @@ apply (subst s) apply (subst r) apply (simp only: power_mult_distrib) - apply (subst mult_assoc [where b=s]) - apply (subst mult_assoc [where a=u]) - apply (subst mult_assoc [where b=u, symmetric]) + apply (subst mult.assoc [where b=s]) + apply (subst mult.assoc [where a=u]) + apply (subst mult.assoc [where b=u, symmetric]) apply (subst u [symmetric]) apply (simp add: mult_ac power_add [symmetric]) done diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Inner_Product.thy --- a/src/HOL/Library/Inner_Product.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Inner_Product.thy Fri Jul 04 20:18:47 2014 +0200 @@ -136,7 +136,7 @@ by (simp add: real_sqrt_mult_distrib) then show "norm (a *\<^sub>R x) = \a\ * norm x" unfolding norm_eq_sqrt_inner - by (simp add: power2_eq_square mult_assoc) + by (simp add: power2_eq_square mult.assoc) qed end @@ -209,11 +209,11 @@ instance proof fix x y z r :: real show "inner x y = inner y x" - unfolding inner_real_def by (rule mult_commute) + unfolding inner_real_def by (rule mult.commute) show "inner (x + y) z = inner x z + inner y z" unfolding inner_real_def by (rule distrib_right) show "inner (scaleR r x) y = r * inner x y" - unfolding inner_real_def real_scaleR_def by (rule mult_assoc) + unfolding inner_real_def real_scaleR_def by (rule mult.assoc) show "0 \ inner x x" unfolding inner_real_def by simp show "inner x x = 0 \ x = 0" @@ -233,7 +233,7 @@ instance proof fix x y z :: complex and r :: real show "inner x y = inner y x" - unfolding inner_complex_def by (simp add: mult_commute) + unfolding inner_complex_def by (simp add: mult.commute) show "inner (x + y) z = inner x z + inner y z" unfolding inner_complex_def by (simp add: distrib_right) show "inner (scaleR r x) y = r * inner x y" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Lattice_Algebras.thy --- a/src/HOL/Library/Lattice_Algebras.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Lattice_Algebras.thy Fri Jul 04 20:18:47 2014 +0200 @@ -13,7 +13,7 @@ apply (rule antisym) apply (simp_all add: le_infI) apply (rule add_le_imp_le_left [of "uminus a"]) - apply (simp only: add_assoc [symmetric], simp add: diff_le_eq add.commute) + apply (simp only: add.assoc [symmetric], simp add: diff_le_eq add.commute) done lemma add_inf_distrib_right: "inf a b + c = inf (a + c) (b + c)" @@ -21,7 +21,7 @@ have "c + inf a b = inf (c + a) (c + b)" by (simp add: add_inf_distrib_left) then show ?thesis - by (simp add: add_commute) + by (simp add: add.commute) qed end @@ -32,10 +32,10 @@ lemma add_sup_distrib_left: "a + sup b c = sup (a + b) (a + c)" apply (rule antisym) apply (rule add_le_imp_le_left [of "uminus a"]) - apply (simp only: add_assoc [symmetric], simp) + apply (simp only: add.assoc [symmetric], simp) apply (simp add: le_diff_eq add.commute) apply (rule le_supI) - apply (rule add_le_imp_le_left [of "a"], simp only: add_assoc[symmetric], simp)+ + apply (rule add_le_imp_le_left [of "a"], simp only: add.assoc[symmetric], simp)+ done lemma add_sup_distrib_right: "sup a b + c = sup (a + c) (b + c)" @@ -43,7 +43,7 @@ have "c + sup a b = sup (c+a) (c+b)" by (simp add: add_sup_distrib_left) then show ?thesis - by (simp add: add_commute) + by (simp add: add.commute) qed end @@ -151,7 +151,7 @@ then show ?r apply - apply (rule add_le_imp_le_right[of _ "uminus b" _]) - apply (simp add: add_assoc) + apply (simp add: add.assoc) done next assume ?r @@ -243,7 +243,7 @@ then have "a + a + - a = - a" by simp then have "a + (a + - a) = - a" - by (simp only: add_assoc) + by (simp only: add.assoc) then have a: "- a = a" by simp show "a = 0" @@ -309,7 +309,7 @@ proof - from add_le_cancel_left [of "uminus a" "plus a a" zero] have "a \ - a \ a + a \ 0" - by (simp add: add_assoc[symmetric]) + by (simp add: add.assoc[symmetric]) then show ?thesis by simp qed @@ -318,7 +318,7 @@ proof - from add_le_cancel_left [of "uminus a" zero "plus a a"] have "- a \ a \ 0 \ a + a" - by (simp add: add_assoc[symmetric]) + by (simp add: add.assoc[symmetric]) then show ?thesis by simp qed diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/ListVector.thy --- a/src/HOL/Library/ListVector.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/ListVector.thy Fri Jul 04 20:18:47 2014 +0200 @@ -94,7 +94,7 @@ apply(simp) apply(case_tac zs) apply(simp) -apply(simp add: add_assoc) +apply(simp add: add.assoc) done subsection "Inner product" @@ -146,7 +146,7 @@ apply simp apply(case_tac ys) apply (simp) -apply (simp add: distrib_left mult_assoc) +apply (simp add: distrib_left mult.assoc) done end diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Multiset.thy Fri Jul 04 20:18:47 2014 +0200 @@ -225,8 +225,8 @@ (* shorter: by (simp add: multiset_eq_iff) fastforce *) proof assume ?rhs then show ?lhs - by (auto simp add: add_assoc add_commute [of "{#b#}"]) - (drule sym, simp add: add_assoc [symmetric]) + by (auto simp add: add.assoc add.commute [of "{#b#}"]) + (drule sym, simp add: add.assoc [symmetric]) next assume ?lhs show ?rhs @@ -1494,9 +1494,9 @@ case (Suc i') with Cons show ?thesis apply simp - apply (subst add_assoc) - apply (subst add_commute [of "{#v#}" "{#x#}"]) - apply (subst add_assoc [symmetric]) + apply (subst add.assoc) + apply (subst add.commute [of "{#v#}" "{#x#}"]) + apply (subst add.assoc [symmetric]) apply simp apply (rule mset_le_multiset_union_diff_commute) apply (simp add: mset_le_single nth_mem_multiset_of) @@ -1597,7 +1597,7 @@ with wf_hyp have "\M \ ?W. M + {#x#} \ ?W" by blast moreover from add have "M0 + K \ ?W" by simp ultimately have "(M0 + K) + {#x#} \ ?W" .. - then show "M0 + (K + {#x#}) \ ?W" by (simp only: add_assoc) + then show "M0 + (K + {#x#}) \ ?W" by (simp only: add.assoc) qed then show "N \ ?W" by (simp only: N) qed @@ -1652,7 +1652,7 @@ apply (rule_tac x = I in exI) apply (simp (no_asm)) apply (rule_tac x = "(K - {#a#}) + Ka" in exI) - apply (simp (no_asm_simp) add: add_assoc [symmetric]) + apply (simp (no_asm_simp) add: add.assoc [symmetric]) apply (drule_tac f = "\M. M - {#a#}" and x="?S + ?T" in arg_cong) apply (simp add: diff_union_single_conv) apply (simp (no_asm_use) add: trans_def) @@ -1695,7 +1695,7 @@ (I + {# x :# K. (x, a) \ r #}) + J') \ mult r") prefer 2 apply force -apply (simp (no_asm_use) add: add_assoc [symmetric] mult_def) +apply (simp (no_asm_use) add: add.assoc [symmetric] mult_def) apply (erule trancl_trans) apply (rule r_into_trancl) apply (simp add: mult1_def set_of_def) @@ -1760,7 +1760,7 @@ apply auto apply (rule_tac x = a in exI) apply (rule_tac x = "C + M0" in exI) -apply (simp add: add_assoc) +apply (simp add: add.assoc) done lemma union_less_mono2: "B \# D ==> C + B \# C + (D::'a::order multiset)" @@ -1771,8 +1771,8 @@ done lemma union_less_mono1: "B \# D ==> B + C \# D + (C::'a::order multiset)" -apply (subst add_commute [of B C]) -apply (subst add_commute [of D C]) +apply (subst add.commute [of B C]) +apply (subst add.commute [of D C]) apply (erule union_less_mono2) done @@ -1941,13 +1941,13 @@ lemmas multi_count_eq = multiset_eq_iff [symmetric] lemma union_commute: "M + N = N + (M::'a multiset)" - by (fact add_commute) + by (fact add.commute) lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" - by (fact add_assoc) + by (fact add.assoc) lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" - by (fact add_left_commute) + by (fact add.left_commute) lemmas union_ac = union_assoc union_commute union_lcomm diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Nat_Bijection.thy --- a/src/HOL/Library/Nat_Bijection.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Nat_Bijection.thy Fri Jul 04 20:18:47 2014 +0200 @@ -61,7 +61,7 @@ "prod_decode (triangle k + m) = prod_decode_aux k m" apply (induct k arbitrary: m) apply (simp add: prod_decode_def) -apply (simp only: triangle_Suc add_assoc) +apply (simp only: triangle_Suc add.assoc) apply (subst prod_decode_aux.simps, simp) done @@ -304,7 +304,7 @@ apply (erule finite_induct, simp) apply (case_tac x) apply (simp add: even_nat_Suc_div_2 even_set_encode_iff vimage_Suc_insert_0) -apply (simp add: finite_vimageI add_commute vimage_Suc_insert_Suc) +apply (simp add: finite_vimageI add.commute vimage_Suc_insert_Suc) apply (simp add: set_encode_def finite_vimage_Suc_iff) done @@ -334,7 +334,7 @@ "n \ set_decode z \ set_decode (2 ^ n + z) = insert n (set_decode z)" apply (induct n arbitrary: z, simp_all) apply (rule set_eqI, induct_tac x, simp, simp add: even_nat_Suc_div_2) - apply (rule set_eqI, induct_tac x, simp, simp add: add_commute) + apply (rule set_eqI, induct_tac x, simp, simp add: add.commute) done lemma finite_set_decode [simp]: "finite (set_decode n)" @@ -377,7 +377,7 @@ by (metis finite_set_decode set_decode_inverse) thus ?thesis using assms apply auto - apply (simp add: set_encode_def nat_add_commute setsum.subset_diff) + apply (simp add: set_encode_def add.commute setsum.subset_diff) done qed thus ?thesis diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Polynomial.thy --- a/src/HOL/Library/Polynomial.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Polynomial.thy Fri Jul 04 20:18:47 2014 +0200 @@ -624,9 +624,9 @@ instance proof fix p q r :: "'a poly" show "(p + q) + r = p + (q + r)" - by (simp add: poly_eq_iff add_assoc) + by (simp add: poly_eq_iff add.assoc) show "p + q = q + p" - by (simp add: poly_eq_iff add_commute) + by (simp add: poly_eq_iff add.commute) show "0 + p = p" by (simp add: poly_eq_iff) qed @@ -713,7 +713,7 @@ lemma degree_add_eq_left: "degree q < degree p \ degree (p + q) = degree p" using degree_add_eq_right [of q p] - by (simp add: add_commute) + by (simp add: add.commute) lemma degree_minus [simp]: "degree (- p) = degree p" unfolding degree_def by simp @@ -824,7 +824,7 @@ by (rule degree_le, simp add: coeff_eq_0) lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p" - by (rule poly_eqI, simp add: mult_assoc) + by (rule poly_eqI, simp add: mult.assoc) lemma smult_0_right [simp]: "smult a 0 = 0" by (rule poly_eqI, simp) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Product_plus.thy --- a/src/HOL/Library/Product_plus.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Product_plus.thy Fri Jul 04 20:18:47 2014 +0200 @@ -81,10 +81,10 @@ subsection {* Class instances *} instance prod :: (semigroup_add, semigroup_add) semigroup_add - by default (simp add: prod_eq_iff add_assoc) + by default (simp add: prod_eq_iff add.assoc) instance prod :: (ab_semigroup_add, ab_semigroup_add) ab_semigroup_add - by default (simp add: prod_eq_iff add_commute) + by default (simp add: prod_eq_iff add.commute) instance prod :: (monoid_add, monoid_add) monoid_add by default (simp_all add: prod_eq_iff) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/RBT_Impl.thy --- a/src/HOL/Library/RBT_Impl.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/RBT_Impl.thy Fri Jul 04 20:18:47 2014 +0200 @@ -1416,7 +1416,7 @@ moreover note feven[unfolded feven_def] (* FIXME: why does by(rule feven[unfolded feven_def]) not work? *) ultimately have "P (2 * (n div 2)) kvs" by - - thus ?thesis using True by (metis div_mod_equality' minus_nat.diff_0 nat_mult_commute) + thus ?thesis using True by (metis div_mod_equality' minus_nat.diff_0 mult.commute) next case False note ge0 moreover from "1.prems" have n2: "n div 2 \ length kvs" by simp @@ -1462,7 +1462,7 @@ moreover note geven[unfolded geven_def] ultimately have "Q (2 * (n div 2)) kvs" by - thus ?thesis using True - by(metis div_mod_equality' minus_nat.diff_0 nat_mult_commute) + by(metis div_mod_equality' minus_nat.diff_0 mult.commute) next case False note ge0 moreover from "2.prems" have n2: "n div 2 \ length kvs" by simp diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Saturated.thy --- a/src/HOL/Library/Saturated.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Saturated.thy Fri Jul 04 20:18:47 2014 +0200 @@ -61,7 +61,7 @@ less_sat_def: "x < y \ nat_of x < nat_of y" instance -by default (auto simp add: less_eq_sat_def less_sat_def not_le sat_eq_iff min.coboundedI1 nat_mult_commute) +by default (auto simp add: less_eq_sat_def less_sat_def not_le sat_eq_iff min.coboundedI1 mult.commute) end @@ -117,14 +117,14 @@ case True thus ?thesis by (simp add: sat_eq_iff) next case False with `a \ 0` show ?thesis - by (simp add: sat_eq_iff nat_mult_min_left nat_mult_min_right mult_assoc min.assoc min.absorb2) + by (simp add: sat_eq_iff nat_mult_min_left nat_mult_min_right mult.assoc min.assoc min.absorb2) qed qed next fix a :: "('a::len) sat" show "1 * a = a" apply (simp add: sat_eq_iff) - apply (metis One_nat_def len_gt_0 less_Suc0 less_zeroE linorder_not_less min.absorb_iff1 min_nat_of_len_of nat_mult_1_right nat_mult_commute) + apply (metis One_nat_def len_gt_0 less_Suc0 less_zeroE linorder_not_less min.absorb_iff1 min_nat_of_len_of nat_mult_1_right mult.commute) done next fix a b c :: "('a::len) sat" @@ -143,7 +143,7 @@ begin instance -by default (auto simp add: less_eq_sat_def less_sat_def not_le sat_eq_iff min.coboundedI1 nat_mult_commute) +by default (auto simp add: less_eq_sat_def less_sat_def not_le sat_eq_iff min.coboundedI1 mult.commute) end diff -r de51a86fc903 -r cc97b347b301 src/HOL/Library/Set_Algebras.thy --- a/src/HOL/Library/Set_Algebras.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Library/Set_Algebras.thy Fri Jul 04 20:18:47 2014 +0200 @@ -108,7 +108,7 @@ done lemma set_plus_rearrange2: "(a::'a::semigroup_add) +o (b +o C) = (a + b) +o C" - by (auto simp add: elt_set_plus_def add_assoc) + by (auto simp add: elt_set_plus_def add.assoc) lemma set_plus_rearrange3: "((a::'a::semigroup_add) +o B) + C = a +o (B + C)" apply (auto simp add: elt_set_plus_def set_plus_def) @@ -216,7 +216,7 @@ lemma set_times_rearrange2: "(a::'a::semigroup_mult) *o (b *o C) = (a * b) *o C" - by (auto simp add: elt_set_times_def mult_assoc) + by (auto simp add: elt_set_times_def mult.assoc) lemma set_times_rearrange3: "((a::'a::semigroup_mult) *o B) * C = a *o (B * C)" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Limits.thy --- a/src/HOL/Limits.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Limits.thy Fri Jul 04 20:18:47 2014 +0200 @@ -702,7 +702,7 @@ by (intro mult_mono' order_refl norm_g norm_ge_zero mult_nonneg_nonneg K elim) also have "\ = norm (f x) * (B * K)" - by (rule mult_assoc) + by (rule mult.assoc) finally show "norm (f x ** g x) \ norm (f x) * (B * K)" . qed with f show ?thesis @@ -716,7 +716,7 @@ apply (rule add_left) apply (rule scaleR_right) apply (rule scaleR_left) - apply (subst mult_commute) + apply (subst mult.commute) using bounded by fast lemma (in bounded_bilinear) Bfun_prod_Zfun: @@ -1273,7 +1273,7 @@ lemma LIMSEQ_linear: "\ X ----> x ; l > 0 \ \ (\ n. X (n * l)) ----> x" unfolding tendsto_def eventually_sequentially - by (metis div_le_dividend div_mult_self1_is_m le_trans nat_mult_commute) + by (metis div_le_dividend div_mult_self1_is_m le_trans mult.commute) lemma Bseq_inverse_lemma: fixes x :: "'a::real_normed_div_algebra" @@ -1530,7 +1530,7 @@ lemma LIM_offset_zero: fixes a :: "'a::real_normed_vector" shows "f -- a --> L \ (\h. f (a + h)) -- 0 --> L" -by (drule_tac k="a" in LIM_offset, simp add: add_commute) +by (drule_tac k="a" in LIM_offset, simp add: add.commute) lemma LIM_offset_zero_cancel: fixes a :: "'a::real_normed_vector" diff -r de51a86fc903 -r cc97b347b301 src/HOL/List.thy --- a/src/HOL/List.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/List.thy Fri Jul 04 20:18:47 2014 +0200 @@ -3846,7 +3846,7 @@ lemma append_replicate_commute: "replicate n x @ replicate k x = replicate k x @ replicate n x" apply (simp add: replicate_add [THEN sym]) -apply (simp add: add_commute) +apply (simp add: add.commute) done text{* Courtesy of Andreas Lochbihler: *} @@ -4192,14 +4192,14 @@ lemma sublist_shift_lemma: "map fst [p<-zip xs [i.. = (Y * A) * X" by (simp add: mult_assoc) + also have "\ = (Y * A) * X" by (simp add: mult.assoc) also have "\ = X" apply (insert assms) apply (frule left_inverse_matrix_dim) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Metis_Examples/Big_O.thy --- a/src/HOL/Metis_Examples/Big_O.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Metis_Examples/Big_O.thy Fri Jul 04 20:18:47 2014 +0200 @@ -309,7 +309,7 @@ apply (erule_tac x = x in allE)+ apply (subgoal_tac "c * ca * abs (f x * g x) = (c * abs(f x)) * (ca * abs (g x))") apply (metis (no_types) abs_ge_zero abs_mult mult_mono') -by (metis mult_assoc mult_left_commute abs_of_pos mult_left_commute abs_mult) +by (metis mult.assoc mult.left_commute abs_of_pos mult.left_commute abs_mult) lemma bigo_mult2 [intro]: "f *o O(g) <= O(f * g)" by (metis bigo_mult bigo_refl set_times_mono3 subset_trans) @@ -421,7 +421,7 @@ lemma bigo_const_mult3: "(c\'a\linordered_field) ~= 0 \ f : O(\x. c * f x)" apply (simp add: bigo_def) -by (metis (no_types) abs_mult mult_assoc mult_1 order_refl left_inverse) +by (metis (no_types) abs_mult mult.assoc mult_1 order_refl left_inverse) lemma bigo_const_mult4: "(c\'a\linordered_field) \ 0 \ O(f) \ O(\x. c * f x)" @@ -444,7 +444,7 @@ apply (rule_tac [2] ext) prefer 2 apply simp - apply (simp add: mult_assoc [symmetric] abs_mult) + apply (simp add: mult.assoc [symmetric] abs_mult) (* couldn't get this proof without the step above *) proof - fix g :: "'b \ 'a" and d :: 'a @@ -530,7 +530,7 @@ apply (rule abs_ge_zero) apply (unfold bigo_def) apply (auto simp add: abs_mult) -by (metis abs_ge_zero mult_left_commute mult_left_mono) +by (metis abs_ge_zero mult.left_commute mult_left_mono) lemma bigo_setsum4: "f =o g +o O(h) \ (\x. SUM y : A x. l x y * f(k x y)) =o @@ -710,6 +710,6 @@ lemma bigo_lesso5: "f \C. \x. f x <= g x + C * abs (h x)" apply (simp only: lesso_def bigo_alt_def) apply clarsimp -by (metis add_commute diff_le_eq) +by (metis add.commute diff_le_eq) end diff -r de51a86fc903 -r cc97b347b301 src/HOL/Metis_Examples/Binary_Tree.thy --- a/src/HOL/Metis_Examples/Binary_Tree.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Metis_Examples/Binary_Tree.thy Fri Jul 04 20:18:47 2014 +0200 @@ -66,7 +66,7 @@ qed next case (Br a t1 t2) thus ?case - by (metis n_leaves.simps(2) nat_add_commute reflect.simps(2)) + by (metis n_leaves.simps(2) add.commute reflect.simps(2)) qed lemma n_nodes_reflect: "n_nodes (reflect t) = n_nodes t" @@ -74,7 +74,7 @@ case Lf thus ?case by (metis reflect.simps(1)) next case (Br a t1 t2) thus ?case - by (metis add_commute n_nodes.simps(2) reflect.simps(2)) + by (metis add.commute n_nodes.simps(2) reflect.simps(2)) qed lemma depth_reflect: "depth (reflect t) = depth t" diff -r de51a86fc903 -r cc97b347b301 src/HOL/MicroJava/Comp/CorrComp.thy --- a/src/HOL/MicroJava/Comp/CorrComp.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/MicroJava/Comp/CorrComp.thy Fri Jul 04 20:18:47 2014 +0200 @@ -137,7 +137,7 @@ apply (simp only: append_assoc) apply (erule thin_rl, erule thin_rl) apply (drule_tac x="pre @ instrs0" in spec) -apply (simp add: add_assoc) +apply (simp add: add.assoc) done lemma progression_refl: @@ -196,7 +196,7 @@ apply simp apply (erule thin_rl, erule thin_rl) apply (drule_tac x="pre @ instr # instrs0" in spec) -apply (simp add: add_assoc) +apply (simp add: add.assoc) done diff -r de51a86fc903 -r cc97b347b301 src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Fri Jul 04 20:18:47 2014 +0200 @@ -183,13 +183,13 @@ subsection {* Some frequently useful arithmetic lemmas over vectors. *} instance vec :: (semigroup_mult, finite) semigroup_mult - by default (vector mult_assoc) + by default (vector mult.assoc) instance vec :: (monoid_mult, finite) monoid_mult by default vector+ instance vec :: (ab_semigroup_mult, finite) ab_semigroup_mult - by default (vector mult_commute) + by default (vector mult.commute) instance vec :: (comm_monoid_mult, finite) comm_monoid_mult by default vector @@ -254,7 +254,7 @@ instance vec :: (ring_char_0, finite) ring_char_0 .. lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x" - by (vector mult_assoc) + by (vector mult.assoc) lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x" by (vector field_simps) lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y" @@ -380,14 +380,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_right_distrib setsum_left_distrib 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_right_distrib setsum_left_distrib mult.assoc) apply (subst setsum.commute) apply simp done @@ -399,7 +399,7 @@ lemma matrix_transpose_mul: "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)" - by (simp add: matrix_matrix_mult_def transpose_def vec_eq_iff mult_commute) + by (simp add: matrix_matrix_mult_def transpose_def vec_eq_iff mult.commute) lemma matrix_eq: fixes A B :: "'a::semiring_1 ^ 'n ^ 'm" @@ -452,7 +452,7 @@ lemma matrix_mult_vsum: "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\i. (x$i) *s column i A) (UNIV:: 'n set)" - by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult_commute) + by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult.commute) lemma vector_componentwise: "(x::'a::ring_1^'n) = (\ j. \i\UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)" @@ -498,7 +498,7 @@ lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::real ^ 'n)" - apply (simp add: matrix_def matrix_vector_mult_def vec_eq_iff mult_commute) + apply (simp add: matrix_def matrix_vector_mult_def vec_eq_iff mult.commute) apply clarify apply (rule linear_componentwise[OF lf, symmetric]) done @@ -518,7 +518,7 @@ lemma matrix_vector_column: "(A::'a::comm_semiring_1^'n^_) *v x = setsum (\i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)" - by (simp add: matrix_vector_mult_def transpose_def vec_eq_iff mult_commute) + by (simp add: matrix_vector_mult_def transpose_def vec_eq_iff mult.commute) lemma adjoint_matrix: "adjoint(\x. (A::real^'n^'m) *v x) = (\x. transpose A *v x)" apply (rule adjoint_unique) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Fri Jul 04 20:18:47 2014 +0200 @@ -3283,7 +3283,7 @@ by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and `e > 0` - by (auto simp add:pos_divide_less_eq[OF `e > 0`] mult_commute) + by (auto simp add:pos_divide_less_eq[OF `e > 0`] mult.commute) finally have "y \ S" apply (subst *) apply (rule assms(1)[unfolded convex_alt,rule_format]) @@ -5923,7 +5923,7 @@ by (auto simp: inner_simps) then have "1 \ inverse d * (x \ i - y \ i)" "1 \ inverse d * (y \ i - x \ i)" apply (rule_tac[!] mult_left_le_imp_le[OF _ assms]) - unfolding mult_assoc[symmetric] + unfolding mult.assoc[symmetric] using assms by (auto simp add: field_simps) then have "inverse d * (x \ i * 2) \ 2 + inverse d * (y \ i * 2)" @@ -6493,7 +6493,7 @@ by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and `e > 0` - by (auto simp add:pos_divide_less_eq[OF `e > 0`] mult_commute) + by (auto simp add:pos_divide_less_eq[OF `e > 0`] mult.commute) finally show "y \ s" apply (subst *) apply (rule assms(1)[unfolded convex_alt,rule_format]) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Fri Jul 04 20:18:47 2014 +0200 @@ -127,7 +127,7 @@ "(f has_field_derivative D) (at a within s) \ ((\z. (f z - f a) / (z - a)) ---> D) (at a within s)" proof - have 1: "\w y. ~(w = a) ==> y / (w - a) - D = (y - (w - a)*D)/(w - a)" - by (metis divide_diff_eq_iff eq_iff_diff_eq_0 mult_commute) + by (metis divide_diff_eq_iff eq_iff_diff_eq_0 mult.commute) show ?thesis apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right) apply (simp add: LIM_zero_iff [where l = D, symmetric]) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Multivariate_Analysis/Determinants.thy --- a/src/HOL/Multivariate_Analysis/Determinants.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Multivariate_Analysis/Determinants.thy Fri Jul 04 20:18:47 2014 +0200 @@ -98,7 +98,7 @@ lemma trace_mul_sym: "trace ((A::'a::comm_semiring_1^'n^'m) ** B) = trace (B**A)" apply (simp add: trace_def matrix_matrix_mult_def) apply (subst setsum.commute) - apply (simp add: mult_commute) + apply (simp add: mult.commute) done text {* Definition of determinant. *} @@ -294,7 +294,7 @@ fixes A :: "'a::comm_ring_1^'n^'n" assumes p: "p permutes (UNIV :: 'n::finite set)" shows "det (\ 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_right_distrib mult.assoc[symmetric]) apply (subst sum_permutations_compose_right[OF p]) proof (rule setsum.cong) let ?U = "UNIV :: 'n set" @@ -318,7 +318,7 @@ by blast show "of_int (sign (q \ p)) * setprod (\i. A$ p i$ (q \ p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\i. A$i$q i) ?U" - by (simp only: thp sign_compose[OF qp pp] mult_commute of_int_mult) + by (simp only: thp sign_compose[OF qp pp] mult.commute of_int_mult) qed rule lemma det_permute_columns: @@ -819,7 +819,7 @@ unfolding permutation_permutes by auto have th00: "of_int (sign p) * of_int (sign p) = (1::'a)" "\a. of_int (sign p) * (of_int (sign p) * a) = a" - unfolding mult_assoc[symmetric] + unfolding mult.assoc[symmetric] unfolding of_int_mult[symmetric] by (simp_all add: sign_idempotent) have ths: "?s q = ?s p * ?s (q \ inv p)" @@ -891,7 +891,7 @@ by blast have *: "\(a::real^'n) b. a + b = 0 \ -a = b" apply (drule_tac f="op + (- a)" in cong[OF refl]) - apply (simp only: ab_left_minus add_assoc[symmetric]) + apply (simp only: ab_left_minus add.assoc[symmetric]) apply simp done from c ci @@ -958,7 +958,7 @@ apply (subst U) unfolding setsum.insert[OF fUk kUk] apply (subst th00) - unfolding add_assoc + unfolding add.assoc apply (subst det_row_add) unfolding thd0 unfolding det_row_mul @@ -1278,10 +1278,10 @@ by (fact setprod_singleton_nat_seg) lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2" - by (simp add: eval_nat_numeral setprod_numseg mult_commute) + by (simp add: eval_nat_numeral setprod_numseg mult.commute) lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3" - by (simp add: eval_nat_numeral setprod_numseg mult_commute) + by (simp add: eval_nat_numeral setprod_numseg mult.commute) lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1" by (simp add: det_def sign_id) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy --- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Fri Jul 04 20:18:47 2014 +0200 @@ -95,10 +95,10 @@ unfolding uminus_vec_def by simp instance vec :: (semigroup_add, finite) semigroup_add - by default (simp add: vec_eq_iff add_assoc) + by default (simp add: vec_eq_iff add.assoc) instance vec :: (ab_semigroup_add, finite) ab_semigroup_add - by default (simp add: vec_eq_iff add_commute) + by default (simp add: vec_eq_iff add.commute) instance vec :: (monoid_add, finite) monoid_add by default (simp_all add: vec_eq_iff) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Fri Jul 04 20:18:47 2014 +0200 @@ -4926,7 +4926,7 @@ apply (rule order_trans[OF mult_left_mono]) apply (rule assms) apply (rule setsum_abs_ge_zero) - apply (subst mult_commute) + apply (subst mult.commute) apply (rule mult_left_mono) apply (rule order_trans[of _ "setsum content p"]) apply (rule eq_refl) @@ -4961,7 +4961,7 @@ apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e]) defer unfolding setsum_left_distrib[symmetric] - apply (subst mult_commute) + apply (subst mult.commute) apply (rule mult_left_mono) apply (rule order_trans[of _ "setsum (content \ snd) p"]) apply (rule eq_refl) @@ -6072,7 +6072,7 @@ proof - case goal1 then show ?case - apply (subst mult_commute, subst pos_le_divide_eq[symmetric]) + apply (subst mult.commute, subst pos_le_divide_eq[symmetric]) using d(2)[rule_format,of "q i" i] using q[rule_format] apply (auto simp add: field_simps) @@ -8349,7 +8349,7 @@ then show "norm (f c) * norm (c - t) < e / 3" using False apply - - apply (subst mult_commute) + apply (subst mult.commute) apply (subst pos_less_divide_eq[symmetric]) apply auto done @@ -11081,7 +11081,7 @@ show ?case apply (rule le_less_trans[OF norm_triangle_sub[of ig sg]]) apply (subst real_sum_of_halves[of e,symmetric]) - unfolding add_assoc[symmetric] + unfolding add.assoc[symmetric] apply (rule add_le_less_mono) defer apply (subst norm_minus_commute) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Multivariate_Analysis/Linear_Algebra.thy --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Fri Jul 04 20:18:47 2014 +0200 @@ -425,7 +425,7 @@ unfolding linear_setsum[OF lf] by (simp add: linear_cmul[OF lf]) finally show "f x \ y = x \ ?w" - by (simp add: inner_setsum_left inner_setsum_right mult_commute) + by (simp add: inner_setsum_left inner_setsum_right mult.commute) qed then show ?thesis unfolding adjoint_def choice_iff @@ -848,7 +848,7 @@ from h have "(x + y) \ span_induct_alt_help S" apply (induct rule: span_induct_alt_help.induct) apply simp - unfolding add_assoc + unfolding add.assoc apply (rule span_induct_alt_help_S) apply assumption apply simp @@ -1083,7 +1083,7 @@ also have "\ = (\v\S. u v *\<^sub>R v) + c *\<^sub>R x" by (simp add: setsum.remove [OF fS xS] algebra_simps) also have "\ = c*\<^sub>R x + y" - by (simp add: add_commute u) + by (simp add: add.commute u) finally have "setsum (\v. ?u v *\<^sub>R v) ?S = c*\<^sub>R x + y" . then show ?thesis using th0 by blast next @@ -1095,7 +1095,7 @@ apply auto done show ?thesis using fS xS th0 - by (simp add: th00 add_commute cong del: if_weak_cong) + by (simp add: th00 add.commute cong del: if_weak_cong) qed then show "?h (c*\<^sub>R x + y)" by fast @@ -1472,7 +1472,7 @@ from Basis_le_norm[OF i, of x] show "norm (?g i) \ norm (f i) * norm x" unfolding norm_scaleR - apply (subst mult_commute) + apply (subst mult.commute) apply (rule mult_mono) apply (auto simp add: field_simps) done @@ -1494,7 +1494,7 @@ have "\B. \x. norm (f x) \ B * norm x" using `linear f` by (rule linear_bounded) then show "\K. \x. norm (f x) \ norm x * K" - by (simp add: mult_commute) + by (simp add: mult.commute) qed next assume "bounded_linear f" @@ -1511,7 +1511,7 @@ using lf unfolding linear_conv_bounded_linear by (rule bounded_linear.pos_bounded) then show ?thesis - by (simp only: mult_commute) + by (simp only: mult.commute) qed lemma bounded_linearI': diff -r de51a86fc903 -r cc97b347b301 src/HOL/Multivariate_Analysis/Operator_Norm.thy --- a/src/HOL/Multivariate_Analysis/Operator_Norm.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Multivariate_Analysis/Operator_Norm.thy Fri Jul 04 20:18:47 2014 +0200 @@ -48,7 +48,7 @@ using f.nonneg_bounded by auto then have "\x. norm (f x) / norm x \ b" by (clarify, case_tac "x = 0", - simp_all add: f.zero pos_divide_le_eq mult_commute) + simp_all add: f.zero pos_divide_le_eq mult.commute) then have "bdd_above (range (\x. norm (f x) / norm x))" unfolding bdd_above_def by fast with UNIV_I show ?thesis @@ -110,7 +110,7 @@ also have "onorm f * norm (g x) \ onorm f * (onorm g * norm x)" by (rule mult_left_mono [OF onorm [OF g] onorm_pos_le [OF f]]) finally show "norm ((f \ g) x) \ onorm f * onorm g * norm x" - by (simp add: mult_assoc) + by (simp add: mult.assoc) qed lemma onorm_scaleR_lemma: @@ -124,7 +124,7 @@ have "\r\ * norm (f x) \ \r\ * (onorm f * norm x)" by (intro mult_left_mono onorm abs_ge_zero f) then show "norm (r *\<^sub>R f x) \ \r\ * onorm f * norm x" - by (simp only: norm_scaleR mult_assoc) + by (simp only: norm_scaleR mult.assoc) qed lemma onorm_scaleR: @@ -142,7 +142,7 @@ then have "onorm (\x. inverse r *\<^sub>R r *\<^sub>R f x) \ \inverse r\ * onorm (\x. r *\<^sub>R f x)" by (rule onorm_scaleR_lemma) with `r \ 0` show "\r\ * onorm f \ onorm (\x. r *\<^sub>R f x)" - by (simp add: inverse_eq_divide pos_le_divide_eq mult_commute) + by (simp add: inverse_eq_divide pos_le_divide_eq mult.commute) qed qed (simp add: onorm_zero) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Multivariate_Analysis/PolyRoots.thy --- a/src/HOL/Multivariate_Analysis/PolyRoots.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Multivariate_Analysis/PolyRoots.thy Fri Jul 04 20:18:47 2014 +0200 @@ -21,7 +21,7 @@ shows "(\i\n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))" using setsum_gp_basic[of x n] apply (simp add: real_of_nat_def) -by (metis eq_iff_diff_eq_0 mult_commute nonzero_eq_divide_eq) +by (metis eq_iff_diff_eq_0 mult.commute nonzero_eq_divide_eq) lemma setsum_power_shift: fixes x :: "'a::{comm_ring,monoid_mult}" @@ -41,7 +41,7 @@ shows "(1 - x) * (\i=m..n. x^i) = x^m - x^Suc n" proof - have "(1 - x) * (\i=m..n. x^i) = x^m * (1 - x) * (\i\n-m. x^i)" - by (metis ab_semigroup_mult_class.mult_ac(1) assms mult_commute setsum_power_shift) + by (metis ab_semigroup_mult_class.mult_ac(1) assms mult.commute setsum_power_shift) also have "... =x^m * (1 - x^Suc(n-m))" by (metis ab_semigroup_mult_class.mult_ac(1) setsum_gp_basic) also have "... = x^m - x^Suc n" @@ -58,7 +58,7 @@ else (x^m - x^Suc n) / (1 - x))" using setsum_gp_multiplied [of m n x] apply (auto simp: real_of_nat_def) -by (metis eq_iff_diff_eq_0 mult_commute nonzero_divide_eq_eq) +by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq) lemma setsum_gp_offset: fixes x :: "'a::{comm_ring,division_ring_inverse_zero}" @@ -133,7 +133,7 @@ proof (induction n) case 0 show ?case - by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult_commute pos_divide_le_eq assms) + by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult.commute pos_divide_le_eq assms) next case (Suc n) then obtain M where M: "\z. M \ norm z \ norm (\i\n. c i * z^i) \ e * norm z ^ Suc n" .. @@ -144,7 +144,7 @@ then have norm1: "0 < norm z" "M \ norm z" "(e + norm (c (Suc n))) / e \ norm z" by auto then have norm2: "(e + norm (c (Suc n))) \ e * norm z" "(norm z * norm z ^ n) > 0" - apply (metis assms less_divide_eq mult_commute not_le) + apply (metis assms less_divide_eq mult.commute not_le) using norm1 apply (metis mult_pos_pos zero_less_power) done have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) = @@ -162,7 +162,7 @@ qed lemma norm_lemma_xy: "\abs b + 1 \ norm(y) - a; norm(x) \ a\ \ b \ norm(x + y)" - by (metis abs_add_one_not_less_self add_commute diff_le_eq dual_order.trans le_less_linear + by (metis abs_add_one_not_less_self add.commute diff_le_eq dual_order.trans le_less_linear norm_diff_ineq) lemma polyfun_extremal: diff -r de51a86fc903 -r cc97b347b301 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Jul 04 20:18:47 2014 +0200 @@ -2725,7 +2725,7 @@ lemma bounded_any_center: "bounded S \ (\e. \y\S. dist a y \ e)" unfolding bounded_def - by auto (metis add_commute add_le_cancel_right dist_commute dist_triangle_le) + by auto (metis add.commute add_le_cancel_right dist_commute dist_triangle_le) lemma bounded_iff: "bounded S \ (\a. \x\S. norm x \ a)" unfolding bounded_any_center [where a=0] @@ -2859,7 +2859,7 @@ then show ?thesis unfolding bounded_pos apply (rule_tac x="b*B" in exI) - using b B by (auto simp add: mult_commute) + using b B by (auto simp add: mult.commute) qed lemma bounded_scaling: @@ -2901,7 +2901,7 @@ lemma bounded_imp_bdd_below: "bounded S \ bdd_below (S :: real set)" by (auto simp: bounded_def bdd_below_def dist_real_def) - (metis abs_le_D1 add_commute diff_le_eq) + (metis abs_le_D1 add.commute diff_le_eq) (* TODO: remove the following lemmas about Inf and Sup, is now in conditionally complete lattice *) @@ -6801,7 +6801,7 @@ then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" unfolding image_iff Bex_def mem_box apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) - apply (auto simp add: pos_le_divide_eq pos_divide_le_eq mult_commute diff_le_iff inner_distrib inner_diff_left) + apply (auto simp add: pos_le_divide_eq pos_divide_le_eq mult.commute diff_le_iff inner_distrib inner_diff_left) done } moreover @@ -6811,7 +6811,7 @@ then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" unfolding image_iff Bex_def mem_box apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) - apply (auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff inner_distrib inner_diff_left) + apply (auto simp add: neg_le_divide_eq neg_divide_le_eq mult.commute diff_le_iff inner_distrib inner_diff_left) done } ultimately show ?thesis using False by (auto simp: cbox_def) @@ -7175,7 +7175,7 @@ then show "norm (f b) / norm b * norm x \ norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]] unfolding f.scaleR and ba using `x\0` `a\0` - by (auto simp add: mult_commute pos_le_divide_eq pos_divide_le_eq) + by (auto simp add: mult.commute pos_le_divide_eq pos_divide_le_eq) qed } ultimately show ?thesis by auto @@ -7406,7 +7406,7 @@ proof (cases "d = 0") case True have *: "\x. ((1 - c) * x \ 0) = (x \ 0)" using `1 - c > 0` - by (metis mult_zero_left mult_commute real_mult_le_cancel_iff1) + by (metis mult_zero_left mult.commute real_mult_le_cancel_iff1) from True have "\n. z n = z0" using cf_z2[of 0] and c unfolding z_def by (simp add: *) then show ?thesis using `e>0` by auto @@ -7431,12 +7431,12 @@ have "dist (z m) (z n) \ c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" using cf_z2[of n "m - n"] and `m>n` unfolding pos_le_divide_eq[OF `1-c>0`] - by (auto simp add: mult_commute dist_commute) + by (auto simp add: mult.commute dist_commute) also have "\ \ c ^ N * d * (1 - c ^ (m - n)) / (1 - c)" using mult_right_mono[OF * order_less_imp_le[OF **]] - unfolding mult_assoc by auto + unfolding mult.assoc by auto also have "\ < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)" - using mult_strict_right_mono[OF N **] unfolding mult_assoc by auto + using mult_strict_right_mono[OF N **] unfolding mult.assoc by auto also have "\ = e * (1 - c ^ (m - n))" using c and `d>0` and `1 - c > 0` by auto also have "\ \ e" using c and `1 - c ^ (m - n) > 0` and `e>0` diff -r de51a86fc903 -r cc97b347b301 src/HOL/NSA/CLim.thy --- a/src/HOL/NSA/CLim.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/NSA/CLim.thy Fri Jul 04 20:18:47 2014 +0200 @@ -22,7 +22,7 @@ lemma all_shift: "(\x::'a::comm_ring_1. P x) = (\x. P (x-a))"; apply auto apply (drule_tac x="x+a" in spec) -apply (simp add: add_assoc) +apply (simp add: add.assoc) done lemma complex_add_minus_iff [simp]: "(x + - a = (0::complex)) = (x=a)" @@ -142,7 +142,7 @@ apply (drule_tac [2] DERIV_ident [THEN DERIV_mult]) apply (auto simp add: distrib_right real_of_nat_Suc) apply (case_tac "n") -apply (auto simp add: mult_ac add_commute) +apply (auto simp add: mult_ac add.commute) done text{*Nonstandard version*} diff -r de51a86fc903 -r cc97b347b301 src/HOL/NSA/Examples/NSPrimes.thy --- a/src/HOL/NSA/Examples/NSPrimes.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/NSA/Examples/NSPrimes.thy Fri Jul 04 20:18:47 2014 +0200 @@ -274,7 +274,7 @@ apply (subgoal_tac "k \ hypnat_of_nat ` {p. prime p}") apply (force simp add: starprime_def, safe) apply (drule_tac x = x in bspec, auto) -apply (metis add_commute hdvd_diff hdvd_one_eq_one hypnat_diff_add_inverse2 hypnat_one_not_prime) +apply (metis add.commute hdvd_diff hdvd_one_eq_one hypnat_diff_add_inverse2 hypnat_one_not_prime) done end diff -r de51a86fc903 -r cc97b347b301 src/HOL/NSA/HDeriv.thy --- a/src/HOL/NSA/HDeriv.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/NSA/HDeriv.thy Fri Jul 04 20:18:47 2014 +0200 @@ -133,7 +133,7 @@ apply (auto simp add: nsderiv_def) apply (rule ccontr, drule_tac x = h in bspec) apply (drule_tac [2] c = h in approx_mult1) -apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult_assoc) +apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult.assoc) done text{*Differentiability implies continuity @@ -143,15 +143,15 @@ apply (drule approx_minus_iff [THEN iffD1]) apply (drule hypreal_not_eq_minus_iff [THEN iffD1]) apply (drule_tac x = "xa - star_of x" in bspec) - prefer 2 apply (simp add: add_assoc [symmetric]) -apply (auto simp add: mem_infmal_iff [symmetric] add_commute) + prefer 2 apply (simp add: add.assoc [symmetric]) +apply (auto simp add: mem_infmal_iff [symmetric] add.commute) apply (drule_tac c = "xa - star_of x" in approx_mult1) apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] - simp add: mult_assoc nonzero_mult_divide_cancel_right) + simp add: mult.assoc nonzero_mult_divide_cancel_right) apply (drule_tac x3=D in HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult, THEN mem_infmal_iff [THEN iffD1]]) -apply (auto simp add: mult_commute +apply (auto simp add: mult.commute intro: approx_trans approx_minus_iff [THEN iffD2]) done @@ -189,7 +189,7 @@ ==> x - y \ 0" apply (simp add: nonzero_divide_eq_eq) apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add - simp add: mult_assoc mem_infmal_iff [symmetric]) + simp add: mult.assoc mem_infmal_iff [symmetric]) apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) done @@ -206,14 +206,14 @@ apply (drule_tac approx_minus_iff [THEN iffD2, THEN bex_Infinitesimal_iff2 [THEN iffD2]]) apply (auto intro!: approx_add_mono1 - simp add: distrib_right distrib_left mult_commute add_assoc) + simp add: distrib_right distrib_left mult.commute add.assoc) apply (rule_tac b1 = "star_of Db * star_of (f x)" - in add_commute [THEN subst]) + in add.commute [THEN subst]) apply (auto intro!: Infinitesimal_add_approx_self2 [THEN approx_sym] Infinitesimal_add Infinitesimal_mult Infinitesimal_star_of_mult Infinitesimal_star_of_mult2 - simp add: add_assoc [symmetric]) + simp add: add.assoc [symmetric]) done text{*Multiplying by a constant*} @@ -309,7 +309,7 @@ proof - assume z: "z \ 0" have "x * y = x * (inverse z * z) * y" by (simp add: z) - thus ?thesis by (simp add: mult_assoc) + thus ?thesis by (simp add: mult.assoc) qed text{*This proof uses both definitions of differentiability.*} @@ -373,7 +373,7 @@ subsubsection {* Equivalence of NS and Standard definitions *} lemma divideR_eq_divide: "x /\<^sub>R y = x / y" -by (simp add: divide_inverse mult_commute) +by (simp add: divide_inverse mult.commute) text{*Now equivalence between NSDERIV and DERIV*} lemma NSDERIV_DERIV_iff: "(NSDERIV f x :> D) = (DERIV f x :> D)" @@ -403,7 +403,7 @@ lemma CARAT_NSDERIV: "NSDERIV f x :> l ==> \g. (\z. f z - f x = g z * (z-x)) & isNSCont g x & g x = l" by (auto simp add: NSDERIV_DERIV_iff isNSCont_isCont_iff CARAT_DERIV - mult_commute) + mult.commute) lemma hypreal_eq_minus_iff3: "(x = y + z) = (x + -z = (y::hypreal))" by auto diff -r de51a86fc903 -r cc97b347b301 src/HOL/NSA/HLim.thy --- a/src/HOL/NSA/HLim.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/NSA/HLim.thy Fri Jul 04 20:18:47 2014 +0200 @@ -95,7 +95,7 @@ lemma NSLIM_zero_cancel: "(%x. f(x) - l) -- x --NS> 0 ==> f -- x --NS> l" apply (drule_tac g = "%x. l" and m = l in NSLIM_add) -apply (auto simp add: add_assoc) +apply (auto simp add: add.assoc) done lemma NSLIM_const_not_eq: @@ -243,10 +243,10 @@ apply (drule_tac [2] x = "- star_of a + x" in spec, safe, simp) apply (erule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]]) apply (erule_tac [3] approx_minus_iff2 [THEN iffD1]) - prefer 2 apply (simp add: add_commute) + prefer 2 apply (simp add: add.commute) apply (rule_tac x = x in star_cases) apply (rule_tac [2] x = x in star_cases) -apply (auto simp add: starfun star_of_def star_n_minus star_n_add add_assoc star_n_zero_num) +apply (auto simp add: starfun star_of_def star_n_minus star_n_add add.assoc star_n_zero_num) done lemma NSLIM_isCont_iff: "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)" diff -r de51a86fc903 -r cc97b347b301 src/HOL/NSA/HTranscendental.thy --- a/src/HOL/NSA/HTranscendental.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/NSA/HTranscendental.thy Fri Jul 04 20:18:47 2014 +0200 @@ -254,7 +254,7 @@ apply (drule_tac x = x in bspec, auto) apply (drule_tac c = x in approx_mult1) apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] - simp add: mult_assoc) + simp add: mult.assoc) apply (rule approx_add_right_cancel [where d="-1"]) apply (rule approx_sym [THEN [2] approx_trans2]) apply (auto simp add: mem_infmal_iff) @@ -424,7 +424,7 @@ apply (drule bspec [where x = x], auto) apply (drule approx_mult1 [where c = x]) apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] - simp add: mult_assoc) + simp add: mult.assoc) done lemma HFinite_cos [simp]: "sumhr (0, whn, %n. cos_coeff n * x ^ n) \ HFinite" @@ -448,7 +448,7 @@ apply auto apply (drule approx_mult1 [where c = x]) apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] - simp add: mult_assoc) + simp add: mult.assoc) apply (rule approx_add_right_cancel [where d = "-1"]) apply simp done @@ -463,7 +463,7 @@ apply (drule bspec [where x = x], auto) apply (drule approx_mult1 [where c = x]) apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] - simp add: mult_assoc) + simp add: mult.assoc) done lemma STAR_sin_cos_Infinitesimal_mult: @@ -481,7 +481,7 @@ "N \ HNatInfinite ==> hypreal_of_real a = hypreal_of_hypnat N * (inverse(hypreal_of_hypnat N) * hypreal_of_real a)" -by (simp add: mult_assoc [symmetric] zero_less_HNatInfinite) +by (simp add: mult.assoc [symmetric] zero_less_HNatInfinite) lemma STAR_sin_Infinitesimal_divide: "[|x \ Infinitesimal; x \ 0 |] ==> ( *f* sin) x/x @= 1" @@ -535,7 +535,7 @@ ==> hypreal_of_hypnat n * ( *f* sin) (hypreal_of_real pi/(hypreal_of_hypnat n)) @= hypreal_of_real pi" -apply (rule mult_commute [THEN subst]) +apply (rule mult.commute [THEN subst]) apply (erule STAR_sin_pi_divide_HNatInfinite_approx_pi) done @@ -564,7 +564,7 @@ apply (auto simp add: starfun_mult [symmetric] starfunNat_real_of_nat divide_inverse) apply (rule_tac f1 = inverse in starfun_o2 [THEN subst]) apply (auto dest: STAR_sin_pi_divide_HNatInfinite_approx_pi - simp add: starfunNat_real_of_nat mult_commute divide_inverse) + simp add: starfunNat_real_of_nat mult.commute divide_inverse) done lemma NSLIMSEQ_cos_one: "(%n. cos (pi / real n))----NS> 1" @@ -587,7 +587,7 @@ "x \ Infinitesimal ==> ( *f* cos) x @= 1 - x\<^sup>2" apply (rule STAR_cos_Infinitesimal [THEN approx_trans]) apply (auto simp add: Infinitesimal_approx_minus [symmetric] - add_assoc [symmetric] numeral_2_eq_2) + add.assoc [symmetric] numeral_2_eq_2) done lemma STAR_cos_Infinitesimal_approx2: diff -r de51a86fc903 -r cc97b347b301 src/HOL/NSA/HyperNat.thy --- a/src/HOL/NSA/HyperNat.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/NSA/HyperNat.thy Fri Jul 04 20:18:47 2014 +0200 @@ -124,7 +124,7 @@ done lemma hypnat_add_self_not_less: "~ (x + y < (x::hypnat))" -by (simp add: linorder_not_le [symmetric] add_commute [of x]) +by (simp add: linorder_not_le [symmetric] add.commute [of x]) lemma hypnat_diff_split: "P(a - b::hypnat) = ((a P 0) & (ALL d. a = b + d --> P d))" @@ -164,9 +164,9 @@ lemma of_nat_eq_add [rule_format]: "\d::hypnat. of_nat m = of_nat n + d --> d \ range of_nat" apply (induct n) -apply (auto simp add: add_assoc) +apply (auto simp add: add.assoc) apply (case_tac x) -apply (auto simp add: add_commute [of 1]) +apply (auto simp add: add.commute [of 1]) done lemma Nats_diff [simp]: "[|a \ Nats; b \ Nats|] ==> (a-b :: hypnat) \ Nats" diff -r de51a86fc903 -r cc97b347b301 src/HOL/NSA/NSA.thy --- a/src/HOL/NSA/NSA.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/NSA/NSA.thy Fri Jul 04 20:18:47 2014 +0200 @@ -430,7 +430,7 @@ apply (drule_tac x="r / \star_of a\" in bspec) apply (simp add: Reals_eq_Standard) apply simp -apply (simp add: hnorm_scaleR pos_less_divide_eq mult_commute) +apply (simp add: hnorm_scaleR pos_less_divide_eq mult.commute) done lemma Compl_HFinite: "- HFinite = HInfinite" @@ -473,11 +473,11 @@ lemma HInfinite_add_ge_zero: "[|(x::hypreal) \ HInfinite; 0 \ y; 0 \ x|] ==> (x + y): HInfinite" by (auto intro!: hypreal_add_zero_less_le_mono - simp add: abs_if add_commute add_nonneg_nonneg HInfinite_def) + simp add: abs_if add.commute add_nonneg_nonneg HInfinite_def) lemma HInfinite_add_ge_zero2: "[|(x::hypreal) \ HInfinite; 0 \ y; 0 \ x|] ==> (y + x): HInfinite" -by (auto intro!: HInfinite_add_ge_zero simp add: add_commute) +by (auto intro!: HInfinite_add_ge_zero simp add: add.commute) lemma HInfinite_add_gt_zero: "[|(x::hypreal) \ HInfinite; 0 < y; 0 < x|] ==> (x + y): HInfinite" @@ -621,13 +621,13 @@ by (simp add: approx_def) lemma approx_minus_iff2: " (x @= y) = (-y + x @= 0)" -by (simp add: approx_def add_commute) +by (simp add: approx_def add.commute) lemma approx_refl [iff]: "x @= x" by (simp add: approx_def Infinitesimal_def) lemma hypreal_minus_distrib1: "-(y + -(x::'a::ab_group_add)) = x + -y" -by (simp add: add_commute) +by (simp add: add.commute) lemma approx_sym: "x @= y ==> y @= x" apply (simp add: approx_def) @@ -688,7 +688,7 @@ lemma approx_minus: "a @= b ==> -a @= -b" apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym]) apply (drule approx_minus_iff [THEN iffD1]) -apply (simp add: add_commute) +apply (simp add: add.commute) done lemma approx_minus2: "-a @= -b ==> a @= b" @@ -746,17 +746,17 @@ lemma Infinitesimal_add_approx: "[| y \ Infinitesimal; x + y = z |] ==> x @= z" apply (rule bex_Infinitesimal_iff [THEN iffD1]) apply (drule Infinitesimal_minus_iff [THEN iffD2]) -apply (auto simp add: add_assoc [symmetric]) +apply (auto simp add: add.assoc [symmetric]) done lemma Infinitesimal_add_approx_self: "y \ Infinitesimal ==> x @= x + y" apply (rule bex_Infinitesimal_iff [THEN iffD1]) apply (drule Infinitesimal_minus_iff [THEN iffD2]) -apply (auto simp add: add_assoc [symmetric]) +apply (auto simp add: add.assoc [symmetric]) done lemma Infinitesimal_add_approx_self2: "y \ Infinitesimal ==> x @= y + x" -by (auto dest: Infinitesimal_add_approx_self simp add: add_commute) +by (auto dest: Infinitesimal_add_approx_self simp add: add.commute) lemma Infinitesimal_add_minus_approx_self: "y \ Infinitesimal ==> x @= x + -y" by (blast intro!: Infinitesimal_add_approx_self Infinitesimal_minus_iff [THEN iffD2]) @@ -770,7 +770,7 @@ "[| y \ Infinitesimal; x @= z + y|] ==> x @= z" apply (drule_tac x = z in Infinitesimal_add_approx_self2 [THEN approx_sym]) apply (erule approx_trans3 [THEN approx_sym]) -apply (simp add: add_commute) +apply (simp add: add.commute) apply (erule approx_sym) done @@ -781,7 +781,7 @@ lemma approx_add_right_cancel: "b + d @= c + d ==> b @= c" apply (rule approx_add_left_cancel) -apply (simp add: add_commute) +apply (simp add: add.commute) done lemma approx_add_mono1: "b @= c ==> d + b @= d + c" @@ -790,19 +790,19 @@ done lemma approx_add_mono2: "b @= c ==> b + a @= c + a" -by (simp add: add_commute approx_add_mono1) +by (simp add: add.commute approx_add_mono1) lemma approx_add_left_iff [simp]: "(a + b @= a + c) = (b @= c)" by (fast elim: approx_add_left_cancel approx_add_mono1) lemma approx_add_right_iff [simp]: "(b + a @= c + a) = (b @= c)" -by (simp add: add_commute) +by (simp add: add.commute) lemma approx_HFinite: "[| x \ HFinite; x @= y |] ==> y \ HFinite" apply (drule bex_Infinitesimal_iff2 [THEN iffD2], safe) apply (drule Infinitesimal_subset_HFinite [THEN subsetD, THEN HFinite_minus_iff [THEN iffD2]]) apply (drule HFinite_add) -apply (auto simp add: add_assoc) +apply (auto simp add: add.assoc) done lemma approx_star_of_HFinite: "x @= star_of D ==> x \ HFinite" @@ -851,7 +851,7 @@ lemma approx_SReal_mult_cancel_zero: "[| (a::hypreal) \ Reals; a \ 0; a*x @= 0 |] ==> x @= 0" apply (drule Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]]) -apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric]) +apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) done lemma approx_mult_SReal1: "[| (a::hypreal) \ Reals; x @= 0 |] ==> x*a @= 0" @@ -867,7 +867,7 @@ lemma approx_SReal_mult_cancel: "[| (a::hypreal) \ Reals; a \ 0; a* w @= a*z |] ==> w @= z" apply (drule Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]]) -apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric]) +apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) done lemma approx_SReal_mult_cancel_iff1 [simp]: @@ -991,7 +991,7 @@ shows "[| y \ 0; y \ Infinitesimal; x/y \ HFinite |] ==> x \ Infinitesimal" apply (drule Infinitesimal_HFinite_mult2, assumption) -apply (simp add: divide_inverse mult_assoc) +apply (simp add: divide_inverse mult.assoc) done lemma Infinitesimal_SReal_divide: @@ -1161,7 +1161,7 @@ lemma lemma_minus_le_zero: "t \ t + -r ==> r \ (0::hypreal)" apply (drule_tac c = "-t" in add_left_mono) -apply (auto simp add: add_assoc [symmetric]) +apply (auto simp add: add.assoc [symmetric]) done lemma lemma_st_part_le2: @@ -1352,7 +1352,7 @@ apply (drule HFinite_inverse2)+ apply (drule approx_mult2, assumption, auto) apply (drule_tac c = "inverse x" in approx_mult1, assumption) -apply (auto intro: approx_sym simp add: mult_assoc) +apply (auto intro: approx_sym simp add: mult.assoc) done (*Used for NSLIM_inverse, NSLIMSEQ_inverse*) @@ -1372,7 +1372,7 @@ shows "[| x \ HFinite - Infinitesimal; h \ Infinitesimal |] ==> inverse(h + x) @= inverse x" -apply (rule add_commute [THEN subst]) +apply (rule add.commute [THEN subst]) apply (blast intro: inverse_add_Infinitesimal_approx) done @@ -1392,7 +1392,7 @@ apply (auto intro: Infinitesimal_mult) apply (rule ccontr, frule Infinitesimal_inverse_HFinite) apply (frule not_Infinitesimal_not_zero) -apply (auto dest: Infinitesimal_HFinite_mult simp add: mult_assoc) +apply (auto dest: Infinitesimal_HFinite_mult simp add: mult.assoc) done declare Infinitesimal_square_iff [symmetric, simp] @@ -1414,7 +1414,7 @@ apply safe apply (frule HFinite_inverse, assumption) apply (drule not_Infinitesimal_not_zero) -apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric]) +apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) done lemma approx_HFinite_mult_cancel_iff1: @@ -1432,7 +1432,7 @@ lemma HInfinite_HFinite_add: "[| x \ HInfinite; y \ HFinite |] ==> x + y \ HInfinite" apply (rule_tac y = "-y" in HInfinite_HFinite_add_cancel) -apply (auto simp add: add_assoc HFinite_minus_iff) +apply (auto simp add: add.assoc HFinite_minus_iff) done lemma HInfinite_ge_HInfinite: @@ -1454,7 +1454,7 @@ apply (frule HFinite_Infinitesimal_not_zero) apply (drule HFinite_not_Infinitesimal_inverse) apply (safe, drule HFinite_mult) -apply (auto simp add: mult_assoc HFinite_HInfinite_iff) +apply (auto simp add: mult.assoc HFinite_HInfinite_iff) done lemma HInfinite_HFinite_not_Infinitesimal_mult2: @@ -1466,7 +1466,7 @@ apply (drule HFinite_not_Infinitesimal_inverse) apply (safe, drule_tac x="inverse y" in HFinite_mult) apply assumption -apply (auto simp add: mult_assoc [symmetric] HFinite_HInfinite_iff) +apply (auto simp add: mult.assoc [symmetric] HFinite_HInfinite_iff) done lemma HInfinite_gt_SReal: @@ -1634,7 +1634,7 @@ lemma Infinitesimal_add_hrabs_hypreal_of_real_less2: "[| x \ Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |] ==> abs (x + hypreal_of_real r) < hypreal_of_real y" -apply (rule add_commute [THEN subst]) +apply (rule add.commute [THEN subst]) apply (erule Infinitesimal_add_hrabs_hypreal_of_real_less, assumption) done @@ -1684,14 +1684,14 @@ lemma Infinitesimal_square_cancel2 [simp]: "(x::hypreal)*x + y*y \ Infinitesimal ==> y*y \ Infinitesimal" apply (rule Infinitesimal_square_cancel) -apply (rule add_commute [THEN subst]) +apply (rule add.commute [THEN subst]) apply (simp (no_asm)) done lemma HFinite_square_cancel2 [simp]: "(x::hypreal)*x + y*y \ HFinite ==> y*y \ HFinite" apply (rule HFinite_square_cancel) -apply (rule add_commute [THEN subst]) +apply (rule add.commute [THEN subst]) apply (simp (no_asm)) done @@ -2143,7 +2143,7 @@ apply (subst pos_less_divide_eq, assumption) apply (subst pos_less_divide_eq) apply (simp add: real_of_nat_Suc_gt_zero) -apply (simp add: mult_commute) +apply (simp add: mult.commute) done lemma finite_inverse_real_of_posnat_gt_real: diff -r de51a86fc903 -r cc97b347b301 src/HOL/NSA/NSCA.thy --- a/src/HOL/NSA/NSCA.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/NSA/NSCA.thy Fri Jul 04 20:18:47 2014 +0200 @@ -137,7 +137,7 @@ lemma approx_SComplex_mult_cancel_zero: "[| a \ SComplex; a \ 0; a*x @= 0 |] ==> x @= 0" apply (drule Standard_inverse [THEN Standard_subset_HFinite [THEN subsetD]]) -apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric]) +apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) done lemma approx_mult_SComplex1: "[| a \ SComplex; x @= 0 |] ==> x*a @= 0" @@ -153,7 +153,7 @@ lemma approx_SComplex_mult_cancel: "[| a \ SComplex; a \ 0; a* w @= a*z |] ==> w @= z" apply (drule Standard_inverse [THEN Standard_subset_HFinite [THEN subsetD]]) -apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric]) +apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) done lemma approx_SComplex_mult_cancel_iff1 [simp]: diff -r de51a86fc903 -r cc97b347b301 src/HOL/NSA/StarDef.thy --- a/src/HOL/NSA/StarDef.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/NSA/StarDef.thy Fri Jul 04 20:18:47 2014 +0200 @@ -766,16 +766,16 @@ subsection {* Ordered group classes *} instance star :: (semigroup_add) semigroup_add -by (intro_classes, transfer, rule add_assoc) +by (intro_classes, transfer, rule add.assoc) instance star :: (ab_semigroup_add) ab_semigroup_add -by (intro_classes, transfer, rule add_commute) +by (intro_classes, transfer, rule add.commute) instance star :: (semigroup_mult) semigroup_mult -by (intro_classes, transfer, rule mult_assoc) +by (intro_classes, transfer, rule mult.assoc) instance star :: (ab_semigroup_mult) ab_semigroup_mult -by (intro_classes, transfer, rule mult_commute) +by (intro_classes, transfer, rule mult.commute) instance star :: (comm_monoid_add) comm_monoid_add by (intro_classes, transfer, rule comm_monoid_add_class.add_0) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Nat.thy --- a/src/HOL/Nat.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Nat.thy Fri Jul 04 20:18:47 2014 +0200 @@ -244,10 +244,10 @@ by (induct m) simp_all lemma mult_Suc_right [simp]: "m * Suc n = m + (m * n)" - by (induct m) (simp_all add: add_left_commute) + by (induct m) (simp_all add: add.left_commute) lemma add_mult_distrib: "(m + n) * k = (m * k) + ((n * k)::nat)" - by (induct m) (simp_all add: add_assoc) + by (induct m) (simp_all add: add.assoc) instance proof fix n m q :: nat @@ -263,20 +263,15 @@ subsubsection {* Addition *} -lemma nat_add_assoc: "(m + n) + k = m + ((n + k)::nat)" - by (rule add_assoc) - -lemma nat_add_commute: "m + n = n + (m::nat)" - by (rule add_commute) +lemma nat_add_left_cancel: + fixes k m n :: nat + shows "k + m = k + n \ m = n" + by (fact add_left_cancel) -lemma nat_add_left_commute: "x + (y + z) = y + ((x + z)::nat)" - by (rule add_left_commute) - -lemma nat_add_left_cancel [simp]: "(k + m = k + n) = (m = (n::nat))" - by (rule add_left_cancel) - -lemma nat_add_right_cancel [simp]: "(m + k = n + k) = (m=(n::nat))" - by (rule add_right_cancel) +lemma nat_add_right_cancel: + fixes k m n :: nat + shows "m + k = n + k \ m = n" + by (fact add_right_cancel) text {* Reasoning about @{text "m + 0 = 0"}, etc. *} @@ -315,31 +310,31 @@ subsubsection {* Difference *} lemma diff_self_eq_0 [simp]: "(m\nat) - m = 0" - by (induct m) simp_all + by (fact diff_cancel) lemma diff_diff_left: "(i::nat) - j - k = i - (j + k)" - by (induct i j rule: diff_induct) simp_all + by (fact diff_diff_add) lemma Suc_diff_diff [simp]: "(Suc m - n) - Suc k = m - n - k" by (simp add: diff_diff_left) lemma diff_commute: "(i::nat) - j - k = i - k - j" - by (simp add: diff_diff_left add_commute) + by (fact diff_right_commute) lemma diff_add_inverse: "(n + m) - n = (m::nat)" - by (induct n) simp_all + by (fact add_diff_cancel_left') lemma diff_add_inverse2: "(m + n) - n = (m::nat)" - by (simp add: diff_add_inverse add_commute [of m n]) + by (fact add_diff_cancel_right') lemma diff_cancel: "(k + m) - (k + n) = m - (n::nat)" - by (induct k) simp_all + by (fact comm_monoid_diff_class.add_diff_cancel_left) lemma diff_cancel2: "(m + k) - (n + k) = m - (n::nat)" - by (simp add: diff_cancel add_commute) + by (fact add_diff_cancel_right) lemma diff_add_0: "n - (n + m) = (0::nat)" - by (induct n) simp_all + by (fact diff_add_zero) lemma diff_Suc_1 [simp]: "Suc n - 1 = n" unfolding One_nat_def by simp @@ -350,20 +345,14 @@ by (induct m n rule: diff_induct) (simp_all add: diff_cancel) lemma diff_mult_distrib2: "k * ((m::nat) - n) = (k * m) - (k * n)" -by (simp add: diff_mult_distrib mult_commute [of k]) +by (simp add: diff_mult_distrib mult.commute [of k]) -- {* NOT added as rewrites, since sometimes they are used from right-to-left *} subsubsection {* Multiplication *} -lemma nat_mult_assoc: "(m * n) * k = m * ((n * k)::nat)" - by (rule mult_assoc) - -lemma nat_mult_commute: "m * n = n * (m::nat)" - by (rule mult_commute) - lemma add_mult_distrib2: "k * (m + n) = (k * m) + ((k * n)::nat)" - by (rule distrib_left) + by (fact distrib_left) lemma mult_is_0 [simp]: "((m::nat) * n = 0) = (m=0 | n=0)" by (induct m) auto @@ -402,7 +391,7 @@ qed lemma mult_cancel2 [simp]: "(m * k = n * k) = (m = n | (k = (0::nat)))" - by (simp add: mult_commute) + by (simp add: mult.commute) lemma Suc_mult_cancel1: "(Suc k * m = Suc k * n) = (m = n)" by (subst mult_cancel1) simp @@ -950,7 +939,7 @@ next case (Suc k) have "0 + i < Suc k + i" by (rule add_less_mono1) simp - hence "i < Suc (i + k)" by (simp add: add_commute) + hence "i < Suc (i + k)" by (simp add: add.commute) from trans[OF this lessI Suc step] show ?case by simp qed @@ -1036,7 +1025,7 @@ by (insert add_right_mono [of 0 m n], simp) lemma le_add1: "n \ ((n + m)::nat)" -by (simp add: add_commute, rule le_add2) +by (simp add: add.commute, rule le_add2) lemma less_add_Suc1: "i < Suc (i + m)" by (rule le_less_trans, rule le_add1, rule lessI) @@ -1071,7 +1060,7 @@ done lemma not_add_less2 [iff]: "~ (j + i < (i::nat))" -by (simp add: add_commute) +by (simp add: add.commute) lemma add_leD1: "m + k \ n ==> m \ (n::nat)" apply (rule order_trans [of _ "m+k"]) @@ -1079,7 +1068,7 @@ done lemma add_leD2: "m + k \ n ==> k \ (n::nat)" -apply (simp add: add_commute) +apply (simp add: add.commute) apply (erule add_leD1) done @@ -1103,7 +1092,7 @@ by (simp add: add_diff_inverse linorder_not_less) lemma le_add_diff_inverse2 [simp]: "n \ m ==> (m - n) + n = (m::nat)" -by (simp add: add_commute) +by (simp add: add.commute) lemma Suc_diff_le: "n \ m ==> Suc m - n = Suc (m - n)" by (induct m n rule: diff_induct) simp_all @@ -1135,7 +1124,7 @@ by (induct j k rule: diff_induct) simp_all lemma diff_add_assoc2: "k \ (j::nat) ==> (j + i) - k = (j - k) + i" -by (simp add: add_commute diff_add_assoc) +by (simp add: add.commute diff_add_assoc) lemma le_imp_diff_is_add: "i \ (j::nat) ==> (j - i = k) = (j = k + i)" by (auto simp add: diff_add_inverse2) @@ -1233,7 +1222,7 @@ done lemma mult_less_cancel1 [simp]: "(k * (m::nat) < k * n) = (0 < k & m < n)" -by (simp add: mult_commute [of k]) +by (simp add: mult.commute [of k]) lemma mult_le_cancel1 [simp]: "(k * (m::nat) \ k * n) = (0 < k --> m \ n)" by (simp add: linorder_not_less [symmetric], auto) @@ -1435,7 +1424,7 @@ by (induct n) simp_all from this [of 0] have "of_nat_aux (\i. i + 1) n 1 = of_nat_aux (\i. i + 1) n 0 + 1" by simp - with Suc show ?case by (simp add: add_commute) + with Suc show ?case by (simp add: add.commute) qed end @@ -1693,13 +1682,13 @@ by arith lemma le_diff_conv2: "k \ j ==> (i \ j-k) = (i+k \ (j::nat))" -by arith + by (fact le_diff_conv2) -- {* FIXME delete *} lemma diff_diff_cancel [simp]: "i \ (n::nat) ==> n - (n - i) = i" by arith lemma le_add_diff: "k \ (n::nat) ==> m \ n + m - k" -by arith + by (fact le_add_diff) -- {* FIXME delete *} (*Replaces the previous diff_less and le_diff_less, which had the stronger second premise n\m*) @@ -1847,7 +1836,7 @@ lemma dvd_antisym: "[| m dvd n; n dvd m |] ==> m = (n::nat)" unfolding dvd_def - by (force dest: mult_eq_self_implies_10 simp add: mult_assoc) + by (force dest: mult_eq_self_implies_10 simp add: mult.assoc) text {* @{term "op dvd"} is a partial order *} @@ -1890,7 +1879,7 @@ done lemma dvd_mult_cancel2: "0 (n*m dvd m) = (n = (1::nat))" - apply (subst mult_commute) + apply (subst mult.commute) apply (erule dvd_mult_cancel1) done @@ -1940,7 +1929,7 @@ fixes m n q :: nat assumes "m dvd q" shows "m dvd n + q \ m dvd n" - using assms by (simp add: dvd_plus_eq_right add_commute [of n]) + using assms by (simp add: dvd_plus_eq_right add.commute [of n]) lemma less_eq_dvd_minus: fixes m n :: nat @@ -1949,7 +1938,7 @@ proof - from assms have "n = m + (n - m)" by simp then obtain q where "n = m + q" .. - then show ?thesis by (simp add: dvd_reduce add_commute [of m]) + then show ?thesis by (simp add: dvd_reduce add.commute [of m]) qed lemma dvd_minus_self: @@ -1966,7 +1955,7 @@ by (auto elim: dvd_plusE) also from assms have "\ \ m dvd r * m + n - q" by simp also from assms have "\ \ m dvd (r * m - q) + n" by simp - also have "\ \ m dvd n + (r * m - q)" by (simp add: add_commute) + also have "\ \ m dvd n + (r * m - q)" by (simp add: add.commute) finally show ?thesis . qed diff -r de51a86fc903 -r cc97b347b301 src/HOL/Nominal/Examples/Standardization.thy --- a/src/HOL/Nominal/Examples/Standardization.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Nominal/Examples/Standardization.thy Fri Jul 04 20:18:47 2014 +0200 @@ -213,7 +213,7 @@ prefer 2 apply (erule allE, erule impE, rule refl, erule spec) apply simp - apply (simp only: foldl_conv_fold add_commute fold_plus_listsum_rev) + apply (simp only: foldl_conv_fold add.commute fold_plus_listsum_rev) apply (fastforce simp add: listsum_map_remove1) apply clarify apply (subgoal_tac "\y::name. y \ (x, u, z)") @@ -235,7 +235,7 @@ prefer 2 apply blast apply simp - apply (simp only: foldl_conv_fold add_commute fold_plus_listsum_rev) + apply (simp only: foldl_conv_fold add.commute fold_plus_listsum_rev) apply (fastforce simp add: listsum_map_remove1) done diff -r de51a86fc903 -r cc97b347b301 src/HOL/NthRoot.thy --- a/src/HOL/NthRoot.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/NthRoot.thy Fri Jul 04 20:18:47 2014 +0200 @@ -222,7 +222,7 @@ simp: sgn_zero_iff sgn_mult power_mult[symmetric] abs_mult power_mult_distrib abs_sgn_eq) lemma real_root_commute: "root m (root n x) = root n (root m x)" - by (simp add: real_root_mult_exp [symmetric] mult_commute) + by (simp add: real_root_mult_exp [symmetric] mult.commute) text {* Monotonicity in first argument *} @@ -432,7 +432,7 @@ from n obtain m where m: "n = 2 * m" unfolding even_mult_two_ex .. from m have "sqrt (2 ^ n) = sqrt ((2 ^ m)\<^sup>2)" - by (simp only: power_mult[symmetric] mult_commute) + by (simp only: power_mult[symmetric] mult.commute) then show ?thesis using m by simp qed @@ -512,7 +512,7 @@ proof (rule right_inverse_eq [THEN iffD1, THEN sym]) show "sqrt x / x \ 0" by (simp add: divide_inverse nneg nz) show "inverse (sqrt x) / (sqrt x / x) = 1" - by (simp add: divide_inverse mult_assoc [symmetric] + by (simp add: divide_inverse mult.assoc [symmetric] power2_eq_square [symmetric] real_inv_sqrt_pow2 pos nz) qed qed @@ -603,7 +603,7 @@ "sqrt ((a + c)\<^sup>2 + (b + d)\<^sup>2) \ sqrt (a\<^sup>2 + b\<^sup>2) + sqrt (c\<^sup>2 + d\<^sup>2)" apply (rule power2_le_imp_le, simp) apply (simp add: power2_sum) -apply (simp only: mult_assoc distrib_left [symmetric]) +apply (simp only: mult.assoc distrib_left [symmetric]) apply (rule mult_left_mono) apply (rule power2_le_imp_le) apply (simp add: power2_sum power_mult_distrib) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Num.thy --- a/src/HOL/Num.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Num.thy Fri Jul 04 20:18:47 2014 +0200 @@ -254,8 +254,8 @@ lemma one_plus_numeral_commute: "1 + numeral x = numeral x + 1" apply (induct x) apply simp - apply (simp add: add_assoc [symmetric], simp add: add_assoc) - apply (simp add: add_assoc [symmetric], simp add: add_assoc) + apply (simp add: add.assoc [symmetric], simp add: add.assoc) + apply (simp add: add.assoc [symmetric], simp add: add.assoc) done lemma numeral_inc: "numeral (inc x) = numeral x + 1" @@ -264,7 +264,7 @@ have "numeral x + (1 + numeral x) + 1 = numeral x + (numeral x + 1) + 1" by (simp only: one_plus_numeral_commute) with Bit1 show ?case - by (simp add: add_assoc) + by (simp add: add.assoc) qed simp_all declare numeral.simps [simp del] @@ -350,7 +350,7 @@ lemma numeral_add: "numeral (m + n) = numeral m + numeral n" by (induct n rule: num_induct) - (simp_all only: numeral_One add_One add_inc numeral_inc add_assoc) + (simp_all only: numeral_One add_One add_inc numeral_inc add.assoc) lemma numeral_plus_numeral: "numeral m + numeral n = numeral (m + n)" by (rule numeral_add [symmetric]) @@ -397,20 +397,20 @@ apply simp apply (rule_tac a=x in add_left_imp_eq) apply (rule_tac a=x in add_right_imp_eq) - apply (simp add: add_assoc) - apply (simp add: add_assoc [symmetric], simp add: add_assoc) + apply (simp add: add.assoc) + apply (simp add: add.assoc [symmetric], simp add: add.assoc) apply (rule_tac a=x in add_left_imp_eq) apply (rule_tac a=x in add_right_imp_eq) - apply (simp add: add_assoc) - apply (simp add: add_assoc, simp add: add_assoc [symmetric]) + apply (simp add: add.assoc) + apply (simp add: add.assoc, simp add: add.assoc [symmetric]) done lemma is_num_add_left_commute: "\is_num x; is_num y\ \ x + (y + z) = y + (x + z)" - by (simp only: add_assoc [symmetric] is_num_add_commute) + by (simp only: add.assoc [symmetric] is_num_add_commute) lemmas is_num_normalize = - add_assoc is_num_add_commute is_num_add_left_commute + add.assoc is_num_add_commute is_num_add_left_commute is_num.intros is_num_numeral minus_add @@ -1172,7 +1172,7 @@ minus_zero left_minus right_minus mult_1_left mult_1_right mult_minus_left mult_minus_right - minus_add_distrib minus_minus mult_assoc + minus_add_distrib minus_minus mult.assoc lemmas of_nat_simps = of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult @@ -1225,20 +1225,20 @@ lemma add_numeral_left [simp]: "numeral v + (numeral w + z) = (numeral(v + w) + z)" - by (simp_all add: add_assoc [symmetric]) + by (simp_all add: add.assoc [symmetric]) lemma add_neg_numeral_left [simp]: "numeral v + (- numeral w + y) = (sub v w + y)" "- numeral v + (numeral w + y) = (sub w v + y)" "- numeral v + (- numeral w + y) = (- numeral(v + w) + y)" - by (simp_all add: add_assoc [symmetric]) + by (simp_all add: add.assoc [symmetric]) lemma mult_numeral_left [simp]: "numeral v * (numeral w * z) = (numeral(v * w) * z :: 'a::semiring_numeral)" "- numeral v * (numeral w * y) = (- numeral(v * w) * y :: 'b::ring_1)" "numeral v * (- numeral w * y) = (- numeral(v * w) * y :: 'b::ring_1)" "- numeral v * (- numeral w * y) = (numeral(v * w) * y :: 'b::ring_1)" - by (simp_all add: mult_assoc [symmetric]) + by (simp_all add: mult.assoc [symmetric]) hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec diff -r de51a86fc903 -r cc97b347b301 src/HOL/Number_Theory/Binomial.thy --- a/src/HOL/Number_Theory/Binomial.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Number_Theory/Binomial.thy Fri Jul 04 20:18:47 2014 +0200 @@ -33,7 +33,7 @@ lemma choose_reduce_nat: "0 < (n::nat) \ 0 < k \ (n choose k) = ((n - 1) choose k) + ((n - 1) choose (k - 1))" - by (metis Suc_diff_1 binomial.simps(2) nat_add_commute neq0_conv) + by (metis Suc_diff_1 binomial.simps(2) add.commute neq0_conv) lemma binomial_eq_0: "n < k \ n choose k = 0" by (induct n arbitrary: k) auto @@ -356,7 +356,7 @@ assumes kn: "k \ n" shows "pochhammer (b - of_nat k + 1) k = ((- 1) ^ k :: 'a::comm_ring_1) * pochhammer (- b) k" unfolding pochhammer_minus[OF kn, where b=b] - unfolding mult_assoc[symmetric] + unfolding mult.assoc[symmetric] unfolding power_add[symmetric] by simp @@ -461,7 +461,7 @@ apply (simp add: pochhammer_Suc_setprod fact_altdef_nat h of_nat_setprod setprod.distrib[symmetric] eq' del: One_nat_def power_Suc) unfolding setprod.union_disjoint[OF th0, unfolded eq3, of "of_nat:: nat \ 'a"] eq[unfolded h] - unfolding mult_assoc[symmetric] + unfolding mult.assoc[symmetric] unfolding setprod.distrib[symmetric] apply simp apply (intro setprod.reindex_bij_witness[where i="op - n" and j="op - n"]) @@ -494,7 +494,7 @@ lemma gbinomial_mult_1': "(a gchoose n) * a = of_nat n * (a gchoose n) + of_nat (Suc n) * (a gchoose (Suc n))" - by (simp add: mult_commute gbinomial_mult_1) + by (simp add: mult.commute gbinomial_mult_1) lemma gbinomial_Suc: "a gchoose (Suc k) = (setprod (\i. a - of_nat i) {0 .. k}) / of_nat (fact (Suc k))" @@ -509,7 +509,7 @@ "((a::'a::field_char_0) gchoose (Suc k)) * (of_nat (fact (Suc k)) :: 'a) = (setprod (\i. a - of_nat i) {0 .. k})" using gbinomial_mult_fact[of k a] - by (subst mult_commute) + by (subst mult.commute) lemma gbinomial_Suc_Suc: @@ -531,8 +531,8 @@ unfolding gbinomial_mult_fact' apply (subst fact_Suc) unfolding of_nat_mult - apply (subst mult_commute) - unfolding mult_assoc + apply (subst mult.commute) + unfolding mult.assoc unfolding gbinomial_mult_fact apply (simp add: field_simps) done @@ -597,7 +597,7 @@ from assms have "n * i \ i * k" by simp then have "n * k - n * i \ n * k - i * k" by arith then have "n * (k - i) \ (n - i) * k" - by (simp add: diff_mult_distrib2 nat_mult_commute) + by (simp add: diff_mult_distrib2 mult.commute) then have "of_nat n * of_nat (k - i) \ of_nat (n - i) * (of_nat k :: 'a)" unfolding of_nat_mult[symmetric] of_nat_le_iff . with assms show "of_nat n / of_nat k \ of_nat (n - i) / (of_nat (k - i) :: 'a)" @@ -634,20 +634,20 @@ also have "... = fact (m+r+k) div (fact r * (fact k * fact m))" apply (subst div_mult_div_if_dvd) apply (auto simp: fact_fact_dvd_fact) - apply (metis ab_semigroup_add_class.add_ac(1) add_commute fact_fact_dvd_fact) + apply (metis ab_semigroup_add_class.add_ac(1) add.commute fact_fact_dvd_fact) done also have "... = (fact (m+r+k) * fact (m+r)) div (fact r * (fact k * fact m) * fact (m+r))" apply (subst div_mult_div_if_dvd [symmetric]) apply (auto simp: fact_fact_dvd_fact) - apply (metis dvd_trans dvd.dual_order.refl fact_fact_dvd_fact mult_dvd_mono mult_left_commute) + apply (metis dvd_trans dvd.dual_order.refl fact_fact_dvd_fact mult_dvd_mono mult.left_commute) done also have "... = (fact (m+r+k) div (fact k * fact (m+r)) * (fact (m+r) div (fact r * fact m)))" apply (subst div_mult_div_if_dvd) apply (auto simp: fact_fact_dvd_fact) - apply(metis mult_left_commute) + apply(metis mult.left_commute) done finally show ?thesis - by (simp add: binomial_altdef_nat mult_commute) + by (simp add: binomial_altdef_nat mult.commute) qed lemma choose_mult: diff -r de51a86fc903 -r cc97b347b301 src/HOL/Number_Theory/Cong.thy --- a/src/HOL/Number_Theory/Cong.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Number_Theory/Cong.thy Fri Jul 04 20:18:47 2014 +0200 @@ -275,11 +275,11 @@ lemma cong_mult_lcancel_nat: "coprime k (m::nat) \ [k * a = k * b ] (mod m) = [a = b] (mod m)" - by (simp add: mult_commute cong_mult_rcancel_nat) + by (simp add: mult.commute cong_mult_rcancel_nat) lemma cong_mult_lcancel_int: "coprime k (m::int) \ [k * a = k * b] (mod m) = [a = b] (mod m)" - by (simp add: mult_commute cong_mult_rcancel_int) + by (simp add: mult.commute cong_mult_rcancel_int) (* was zcong_zgcd_zmult_zmod *) lemma coprime_cong_mult_int: @@ -321,7 +321,7 @@ proof (cases "b \ a") case True then show ?rhs using eqm - by (metis cong_altdef_nat dvd_def le_add_diff_inverse add_0_right mult_0 mult_commute) + by (metis cong_altdef_nat dvd_def le_add_diff_inverse add_0_right mult_0 mult.commute) next case False then show ?rhs using eqm @@ -333,7 +333,7 @@ next assume ?rhs then show ?lhs - by (metis cong_nat_def mod_mult_self2 nat_mult_commute) + by (metis cong_nat_def mod_mult_self2 mult.commute) qed lemma cong_gcd_eq_int: "[(a::int) = b] (mod m) \ gcd a m = gcd b m" @@ -467,18 +467,18 @@ apply force apply (cases "a = 0", simp add: cong_0_1_nat) apply (rule iffI) - apply (metis cong_to_1_nat dvd_def monoid_mult_class.mult.right_neutral mult_commute mult_eq_if) + apply (metis cong_to_1_nat dvd_def monoid_mult_class.mult.right_neutral mult.commute mult_eq_if) apply (metis cong_add_lcancel_0_nat cong_mult_self_nat) done lemma cong_le_nat: "(y::nat) <= x \ [x = y] (mod n) \ (\q. x = q * n + y)" - by (metis cong_altdef_nat Nat.le_imp_diff_is_add dvd_def nat_mult_commute) + by (metis cong_altdef_nat Nat.le_imp_diff_is_add dvd_def mult.commute) lemma cong_solve_nat: "(a::nat) \ 0 \ EX x. [a * x = gcd a n] (mod n)" apply (cases "n = 0") apply force apply (frule bezout_nat [of a n], auto) - by (metis cong_add_rcancel_0_nat cong_mult_self_nat mult_commute) + by (metis cong_add_rcancel_0_nat cong_mult_self_nat mult.commute) lemma cong_solve_int: "(a::int) \ 0 \ EX x. [a * x = gcd a n] (mod n)" apply (cases "n = 0") @@ -487,7 +487,7 @@ apply (rule_tac x = "-1" in exI) apply auto apply (insert bezout_int [of a n], auto) - by (metis cong_iff_lin_int mult_commute) + by (metis cong_iff_lin_int mult.commute) lemma cong_solve_dvd_nat: assumes a: "(a::nat) \ 0" and b: "gcd a n dvd d" @@ -603,9 +603,9 @@ from cong_solve_coprime_nat [OF b] obtain x2 where two: "[m2 * x2 = 1] (mod m1)" by auto have "[m1 * x1 = 0] (mod m1)" - by (subst mult_commute, rule cong_mult_self_nat) + by (subst mult.commute, rule cong_mult_self_nat) moreover have "[m2 * x2 = 0] (mod m2)" - by (subst mult_commute, rule cong_mult_self_nat) + by (subst mult.commute, rule cong_mult_self_nat) moreover note one two ultimately show ?thesis by blast qed @@ -622,9 +622,9 @@ from cong_solve_coprime_int [OF b] obtain x2 where two: "[m2 * x2 = 1] (mod m1)" by auto have "[m1 * x1 = 0] (mod m1)" - by (subst mult_commute, rule cong_mult_self_int) + by (subst mult.commute, rule cong_mult_self_int) moreover have "[m2 * x2 = 0] (mod m2)" - by (subst mult_commute, rule cong_mult_self_int) + by (subst mult.commute, rule cong_mult_self_int) moreover note one two ultimately show ?thesis by blast qed @@ -729,7 +729,7 @@ apply (rule cong_trans_nat) prefer 2 apply (rule `[y = u2] (mod m2)`) - apply (subst mult_commute) + apply (subst mult.commute) apply (rule cong_modulus_mult_nat) apply (rule cong_mod_nat) using nz apply auto @@ -778,7 +778,7 @@ by auto moreover have "[(PROD j : A - {i}. m j) * x = 0] (mod (PROD j : A - {i}. m j))" - by (subst mult_commute, rule cong_mult_self_nat) + by (subst mult.commute, rule cong_mult_self_nat) ultimately show "\a. [a = 1] (mod m i) \ [a = 0] (mod setprod m (A - {i}))" by blast @@ -835,7 +835,7 @@ [x = y] (mod (PROD i:A. m i))" apply (induct set: finite) apply auto - apply (metis coprime_cong_mult_nat nat_mult_commute setprod_coprime_nat) + apply (metis coprime_cong_mult_nat mult.commute setprod_coprime_nat) done lemma chinese_remainder_unique_nat: diff -r de51a86fc903 -r cc97b347b301 src/HOL/Number_Theory/Eratosthenes.thy --- a/src/HOL/Number_Theory/Eratosthenes.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Number_Theory/Eratosthenes.thy Fri Jul 04 20:18:47 2014 +0200 @@ -50,7 +50,7 @@ lemma in_numbers_of_marks_eq: "m \ numbers_of_marks n bs \ m \ {n.. bs ! (m - n)" - by (simp add: numbers_of_marks_def in_set_enumerate_eq image_iff add_commute) + by (simp add: numbers_of_marks_def in_set_enumerate_eq image_iff add.commute) lemma sorted_list_of_set_numbers_of_marks: "sorted_list_of_set (numbers_of_marks n bs) = map fst (filter snd (enumerate n bs))" @@ -141,7 +141,7 @@ by (simp add: div_mod_equality' [symmetric]) } then have "w dvd v + w + r + (w - v mod w) \ w dvd m + w + r + (w - m mod w)" - by (simp add: add_assoc add_left_commute [of m] add_left_commute [of v] + by (simp add: add.assoc add.left_commute [of m] add.left_commute [of v] dvd_plus_eq_left dvd_plus_eq_right) moreover from q have "Suc q = m + w + r" by (simp add: w_def) moreover from q have "Suc (Suc q) = v + w + r" by (simp add: v_def w_def) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Number_Theory/Fib.thy --- a/src/HOL/Number_Theory/Fib.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Number_Theory/Fib.thy Fri Jul 04 20:18:47 2014 +0200 @@ -60,7 +60,7 @@ lemma coprime_fib_Suc_nat: "coprime (fib (n::nat)) (fib (Suc n))" apply (induct n rule: fib.induct) apply auto - apply (metis gcd_add1_nat nat_add_commute) + apply (metis gcd_add1_nat add.commute) done lemma gcd_fib_add: "gcd (fib m) (fib (n + m)) = gcd (fib m) (fib n)" @@ -68,7 +68,7 @@ apply (cases m) apply (auto simp add: fib_add) apply (subst gcd_commute_nat) - apply (subst mult_commute) + apply (subst mult.commute) apply (metis coprime_fib_Suc_nat gcd_add_mult_nat gcd_mult_cancel_nat gcd_nat.commute) done diff -r de51a86fc903 -r cc97b347b301 src/HOL/Number_Theory/Gauss.thy --- a/src/HOL/Number_Theory/Gauss.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Number_Theory/Gauss.thy Fri Jul 04 20:18:47 2014 +0200 @@ -260,11 +260,11 @@ fix y::int and z::int assume "p - (y*a) mod p = (z*a) mod p" then have "[(y*a) mod p + (z*a) mod p = 0] (mod p)" - by (metis add_commute diff_eq_eq dvd_refl cong_int_def dvd_eq_mod_eq_0 mod_0) + by (metis add.commute diff_eq_eq dvd_refl cong_int_def dvd_eq_mod_eq_0 mod_0) moreover have "[y * a = (y*a) mod p] (mod p)" by (metis cong_int_def mod_mod_trivial) ultimately have "[a * (y + z) = 0] (mod p)" - by (metis cong_int_def mod_add_left_eq mod_add_right_eq mult_commute ring_class.ring_distribs(1)) + by (metis cong_int_def mod_add_left_eq mod_add_right_eq mult.commute ring_class.ring_distribs(1)) with p_prime a_nonzero p_a_relprime have a: "[y + z = 0] (mod p)" by (metis cong_prime_prod_zero_int) @@ -319,19 +319,19 @@ subsection {* Gauss' Lemma *} lemma aux: "setprod id A * -1 ^ card E * a ^ card A * -1 ^ card E = setprod id A * a ^ card A" -by (metis (no_types) minus_minus mult_commute mult_left_commute power_minus power_one) +by (metis (no_types) minus_minus mult.commute mult.left_commute power_minus power_one) theorem pre_gauss_lemma: "[a ^ nat((int p - 1) div 2) = (-1) ^ (card E)] (mod p)" proof - have "[setprod id A = setprod id F * setprod id D](mod p)" - by (auto simp add: prod_D_F_eq_prod_A mult_commute cong del:setprod.cong) + by (auto simp add: prod_D_F_eq_prod_A mult.commute cong del:setprod.cong) then have "[setprod id A = ((-1)^(card E) * setprod id E) * setprod id D] (mod p)" apply (rule cong_trans_int) apply (metis cong_scalar_int prod_F_zcong) done then have "[setprod id A = ((-1)^(card E) * setprod id C)] (mod p)" - by (metis C_prod_eq_D_times_E mult_commute mult_left_commute) + by (metis C_prod_eq_D_times_E mult.commute mult.left_commute) then have "[setprod id A = ((-1)^(card E) * setprod id B)] (mod p)" by (rule cong_trans_int) (metis C_B_zcong_prod cong_scalar2_int) then have "[setprod id A = ((-1)^(card E) * @@ -349,7 +349,7 @@ then have "[setprod id A = ((-1)^(card E) * a^(card A) * setprod id A)](mod p)" apply (rule cong_trans_int) - apply (simp add: cong_scalar2_int cong_scalar_int finite_A setprod_constant mult_assoc) + apply (simp add: cong_scalar2_int cong_scalar_int finite_A setprod_constant mult.assoc) done then have a: "[setprod id A * (-1)^(card E) = ((-1)^(card E) * a^(card A) * setprod id A * (-1)^(card E))](mod p)" @@ -357,7 +357,7 @@ then have "[setprod id A * (-1)^(card E) = setprod id A * (-1)^(card E) * a^(card A) * (-1)^(card E)](mod p)" apply (rule cong_trans_int) - apply (simp add: a mult_commute mult_left_commute) + apply (simp add: a mult.commute mult.left_commute) done then have "[setprod id A * (-1)^(card E) = setprod id A * a^(card A)](mod p)" apply (rule cong_trans_int) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Number_Theory/Pocklington.thy --- a/src/HOL/Number_Theory/Pocklington.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Number_Theory/Pocklington.thy Fri Jul 04 20:18:47 2014 +0200 @@ -264,7 +264,7 @@ mod_div_equality[of "(n - 1)" m, symmetric] by (simp add:power_add[symmetric] cong_nat_def th3 del: One_nat_def) have th1: "[a ^ ((n - 1) mod m) = 1] (mod n)" - by (metis cong_mult_rcancel_nat nat_mult_commute th2 yn) + by (metis cong_mult_rcancel_nat mult.commute th2 yn) from m(4)[rule_format, OF th0] nm1 less_trans[OF mod_less_divisor[OF m(1), of "n - 1"] m(2)] th1 have False by blast } @@ -385,7 +385,7 @@ from H have onz: "?o \ 0" by (simp add: ord_eq_0) hence op: "?o > 0" by simp from mod_div_equality[of d "ord n a"] lh - have "[a^(?o*?q + ?r) = 1] (mod n)" by (simp add: cong_nat_def mult_commute) + have "[a^(?o*?q + ?r) = 1] (mod n)" by (simp add: cong_nat_def mult.commute) hence "[(a^?o)^?q * (a^?r) = 1] (mod n)" by (simp add: cong_nat_def power_mult[symmetric] power_add[symmetric]) hence th: "[a^?r = 1] (mod n)" @@ -493,7 +493,7 @@ from H[rule_format, of d] h d have "d = 1" by blast} moreover {assume h: "e\<^sup>2 \ n" - from e have "e dvd n" unfolding dvd_def by (simp add: mult_commute) + from e have "e dvd n" unfolding dvd_def by (simp add: mult.commute) with H[rule_format, of e] h have "e=1" by simp with e have "d = n" by simp} ultimately have "d=1 \ d=n" by blast} @@ -552,7 +552,7 @@ hence "a ^ (r * q) + p * 0 = 1 + p * (l*k)" by (simp add: mult_ac) hence odq: "ord p (a^r) dvd q" unfolding ord_divides[symmetric] power_mult[symmetric] - by (metis an cong_dvd_modulus_nat mult_commute nqr pn) + by (metis an cong_dvd_modulus_nat mult.commute nqr pn) from odq[unfolded dvd_def] obtain d where d: "q = ord p (a^r) * d" by blast {assume d1: "d \ 1" obtain P where P: "prime P" "P dvd d" @@ -564,9 +564,9 @@ by (metis zero_not_prime_nat) from P(2) obtain t where t: "d = P*t" unfolding dvd_def by blast from d s t P0 have s': "ord p (a^r) * t = s" - by (metis mult_commute mult_cancel1 nat_mult_assoc) + by (metis mult.commute mult_cancel1 mult.assoc) have "ord p (a^r) * t*r = r * ord p (a^r) * t" - by (metis mult_assoc mult_commute) + by (metis mult.assoc mult.commute) hence exps: "a^(ord p (a^r) * t*r) = ((a ^ r) ^ ord p (a^r)) ^ t" by (simp only: power_mult) then have th: "[((a ^ r) ^ ord p (a^r)) ^ t= 1] (mod p)" @@ -597,7 +597,7 @@ by (metis p01(1)) from p0 d have "p + q * 0 = 1 + q * d" by simp then show ?thesis - by (metis cong_iff_lin_nat mult_commute) + by (metis cong_iff_lin_nat mult.commute) qed theorem pocklington: @@ -767,7 +767,7 @@ by auto from div_mult1_eq[of r q p] p(2) have eq1: "r* (q div p) = (n - 1) div p" - unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult_commute) + unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult.commute) have ath: "\a (b::nat). a <= b \ a \ 0 ==> 1 <= a \ 1 <= b" by arith {assume "a ^ ((n - 1) div p) mod n = 0" then obtain s where s: "a ^ ((n - 1) div p) = n*s" @@ -779,7 +779,7 @@ with eq0 have "a^ (n - 1) = (n*s)^p" by (simp add: power_mult[symmetric]) hence "1 = (n*s)^(Suc (p - 1)) mod n" using bqn p01 by simp - also have "\ = 0" by (simp add: mult_assoc) + also have "\ = 0" by (simp add: mult.assoc) finally have False by simp } then have th11: "a ^ ((n - 1) div p) mod n \ 0" by auto have th1: "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Number_Theory/Primes.thy --- a/src/HOL/Number_Theory/Primes.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Number_Theory/Primes.thy Fri Jul 04 20:18:47 2014 +0200 @@ -106,7 +106,7 @@ lemma not_prime_eq_prod_nat: "(n::nat) > 1 \ ~ prime n \ EX m k. n = m * k & 1 < m & m < n & 1 < k & k < n" unfolding prime_nat_def dvd_def apply auto - by (metis mult_commute linorder_neq_iff linorder_not_le mult_1 + by (metis mult.commute linorder_neq_iff linorder_not_le mult_1 n_less_n_mult_m one_le_mult_iff less_imp_le_nat) lemma prime_dvd_power_nat: "prime p \ p dvd x^n \ p dvd x" @@ -209,7 +209,7 @@ from coprime_dvd_mult_nat[OF pnb pab] have ?thesis by blast } moreover { assume pb: "p dvd b" - have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute) + have pnba: "p^n dvd b*a" using pab by (simp add: mult.commute) from coprime_common_divisor_nat [OF ab, of p] pb p have "\ p dvd a" by auto with p have "coprime a p" @@ -323,7 +323,7 @@ moreover {assume px: "p dvd y" then obtain d where d: "y = p*d" unfolding dvd_def by blast - from Suc.prems d have "p*d*x = p^Suc k" by (simp add: mult_commute) + from Suc.prems d have "p*d*x = p^Suc k" by (simp add: mult.commute) hence th: "d*x = p^k" using p0 by simp from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "x = p^j" by blast with d have "y = p^Suc i" by simp @@ -355,7 +355,7 @@ shows "d dvd p^k \ (\ i. i \ k \ d = p ^i)" proof assume H: "d dvd p^k" then obtain e where e: "d*e = p^k" - unfolding dvd_def apply (auto simp add: mult_commute) by blast + unfolding dvd_def apply (auto simp add: mult.commute) by blast from prime_power_mult[OF p e] obtain i j where ij: "d = p^i" "e=p^j" by blast from e ij have "p^(i + j) = p^k" by (simp add: power_add) hence "i + j = k" using p prime_gt_1_nat power_inject_exp[of p "i+j" k] by simp diff -r de51a86fc903 -r cc97b347b301 src/HOL/Number_Theory/Residues.thy --- a/src/HOL/Number_Theory/Residues.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Number_Theory/Residues.thy Fri Jul 04 20:18:47 2014 +0200 @@ -65,7 +65,7 @@ apply (rule comm_monoid) apply (unfold R_def residue_ring_def, auto) apply (subst mod_add_eq [symmetric]) - apply (subst mult_commute) + apply (subst mult.commute) apply (subst mod_mult_right_eq [symmetric]) apply (simp add: field_simps) done @@ -105,7 +105,7 @@ apply auto apply (metis invertible_coprime_int) apply (subst (asm) coprime_iff_invertible'_int) - apply (auto simp add: cong_int_def mult_commute) + apply (auto simp add: cong_int_def mult.commute) done lemma res_neg_eq: "\ x = (- x) mod m" @@ -120,7 +120,7 @@ apply auto apply (subgoal_tac "y mod m = - x mod m") apply simp - apply (metis minus_add_cancel mod_mult_self1 mult_commute) + apply (metis minus_add_cancel mod_mult_self1 mult.commute) done lemma finite [iff]: "finite (carrier R)" @@ -159,7 +159,7 @@ apply (insert m_gt_one) apply (induct n) apply (auto simp add: nat_pow_def one_cong) - apply (metis mult_commute mult_cong) + apply (metis mult.commute mult_cong) done lemma neg_cong: "\ (x mod m) = (- x) mod m" @@ -196,7 +196,7 @@ lemma one_eq_neg_one: "\ = \ \ \ m = 2" apply (simp add: res_one_eq res_neg_eq) - apply (metis add_commute add_diff_cancel mod_mod_trivial one_add_one uminus_add_conv_diff + apply (metis add.commute add_diff_cancel mod_mod_trivial one_add_one uminus_add_conv_diff zero_neq_one zmod_zminus1_eq_if) done diff -r de51a86fc903 -r cc97b347b301 src/HOL/Old_Number_Theory/Euler.thy --- a/src/HOL/Old_Number_Theory/Euler.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Old_Number_Theory/Euler.thy Fri Jul 04 20:18:47 2014 +0200 @@ -148,7 +148,7 @@ b = "x * (a * MultInv p x)" and c = "a * (x * MultInv p x)" in zcong_trans, force) apply (frule_tac p = p and x = x in MultInv_prop2, auto) -apply (metis StandardRes_SRStar_prop3 mult_1_right mult_commute zcong_sym zcong_zmult_prop1) +apply (metis StandardRes_SRStar_prop3 mult_1_right mult.commute zcong_sym zcong_zmult_prop1) apply (auto simp add: mult_ac) done diff -r de51a86fc903 -r cc97b347b301 src/HOL/Old_Number_Theory/Factorization.thy --- a/src/HOL/Old_Number_Theory/Factorization.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Old_Number_Theory/Factorization.thy Fri Jul 04 20:18:47 2014 +0200 @@ -68,7 +68,7 @@ lemma prod_append: "prod (xs @ ys) = prod xs * prod ys" apply (induct xs) - apply (simp_all add: mult_assoc) + apply (simp_all add: mult.assoc) done lemma prod_xy_prod: diff -r de51a86fc903 -r cc97b347b301 src/HOL/Old_Number_Theory/Fib.thy --- a/src/HOL/Old_Number_Theory/Fib.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Old_Number_Theory/Fib.thy Fri Jul 04 20:18:47 2014 +0200 @@ -106,7 +106,7 @@ apply (case_tac m) apply simp apply (simp add: fib_add) - apply (simp add: add_commute gcd_non_0 [OF fib_Suc_gr_0]) + apply (simp add: add.commute gcd_non_0 [OF fib_Suc_gr_0]) apply (simp add: gcd_non_0 [OF fib_Suc_gr_0, symmetric]) apply (simp add: gcd_fib_Suc_eq_1 gcd_mult_cancel) done diff -r de51a86fc903 -r cc97b347b301 src/HOL/Old_Number_Theory/Gauss.thy --- a/src/HOL/Old_Number_Theory/Gauss.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Old_Number_Theory/Gauss.thy Fri Jul 04 20:18:47 2014 +0200 @@ -342,7 +342,7 @@ apply (elim zcong_trans) by (simp only: zcong_refl) also have "y * a + ya * a = a * (y + ya)" - by (simp add: distrib_left mult_commute) + by (simp add: distrib_left mult.commute) finally have "[a * (y + ya) = 0] (mod p)" . with p_prime a_nonzero zcong_zprime_prod_zero [of p a "y + ya"] p_a_relprime @@ -428,7 +428,7 @@ also have "setprod uminus E = (setprod id E) * (-1)^(card E)" using finite_E by (induct set: finite) auto then have "setprod uminus E = (-1) ^ (card E) * (setprod id E)" - by (simp add: mult_commute) + by (simp add: mult.commute) with two show ?thesis by simp qed @@ -443,7 +443,7 @@ "[a ^ nat((p - 1) div 2) = (-1) ^ (card E)] (mod p)" proof - have "[setprod id A = setprod id F * setprod id D](mod p)" - by (auto simp add: prod_D_F_eq_prod_A mult_commute cong del:setprod.cong) + by (auto simp add: prod_D_F_eq_prod_A mult.commute cong del:setprod.cong) then have "[setprod id A = ((-1)^(card E) * setprod id E) * setprod id D] (mod p)" apply (rule zcong_trans) @@ -452,7 +452,7 @@ then have "[setprod id A = ((-1)^(card E) * setprod id C)] (mod p)" apply (rule zcong_trans) apply (insert C_prod_eq_D_times_E, erule subst) - apply (subst mult_assoc, auto) + apply (subst mult.assoc, auto) done then have "[setprod id A = ((-1)^(card E) * setprod id B)] (mod p)" apply (rule zcong_trans) @@ -473,7 +473,7 @@ then have "[setprod id A = ((-1)^(card E) * a^(card A) * setprod id A)](mod p)" apply (rule zcong_trans) - apply (simp add: zcong_scalar2 zcong_scalar finite_A setprod_constant mult_assoc) + apply (simp add: zcong_scalar2 zcong_scalar finite_A setprod_constant mult.assoc) done then have a: "[setprod id A * (-1)^(card E) = ((-1)^(card E) * a^(card A) * setprod id A * (-1)^(card E))](mod p)" @@ -481,7 +481,7 @@ then have "[setprod id A * (-1)^(card E) = setprod id A * (-1)^(card E) * a^(card A) * (-1)^(card E)](mod p)" apply (rule zcong_trans) - apply (simp add: a mult_commute mult_left_commute) + apply (simp add: a mult.commute mult.left_commute) done then have "[setprod id A * (-1)^(card E) = setprod id A * a^(card A)](mod p)" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Old_Number_Theory/Int2.thy --- a/src/HOL/Old_Number_Theory/Int2.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Old_Number_Theory/Int2.thy Fri Jul 04 20:18:47 2014 +0200 @@ -269,7 +269,7 @@ lemma aux___3: "[j * k = a] (mod p) ==> [(MultInv p j) * j * k = (MultInv p j) * a] (mod p)" - by (auto simp add: mult_assoc zcong_scalar2) + by (auto simp add: mult.assoc zcong_scalar2) lemma aux___4: "[|2 < p; zprime p; ~([j = 0](mod p)); [(MultInv p j) * j * k = (MultInv p j) * a] (mod p) |] diff -r de51a86fc903 -r cc97b347b301 src/HOL/Old_Number_Theory/IntPrimes.thy --- a/src/HOL/Old_Number_Theory/IntPrimes.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Old_Number_Theory/IntPrimes.thy Fri Jul 04 20:18:47 2014 +0200 @@ -93,7 +93,7 @@ apply (simp add: zgcd_greatest_iff) apply (rule_tac n = k in zrelprime_zdvd_zmult) prefer 2 - apply (simp add: mult_commute) + apply (simp add: mult.commute) apply (metis zgcd_1 zgcd_commute zgcd_left_commute) done @@ -142,7 +142,7 @@ "[a = b] (mod m) ==> [c = d] (mod m) ==> [a * c = b * d] (mod m)" apply (rule_tac b = "b * c" in zcong_trans) apply (unfold zcong_def) - apply (metis right_diff_distrib dvd_mult mult_commute) + apply (metis right_diff_distrib dvd_mult mult.commute) apply (metis right_diff_distrib dvd_mult) done @@ -165,7 +165,7 @@ apply (rule_tac [3] s = "a * a - 1 + p * (1 - a)" in subst) prefer 4 apply (simp add: zdvd_reduce) - apply (simp_all add: left_diff_distrib mult_commute right_diff_distrib) + apply (simp_all add: left_diff_distrib mult.commute right_diff_distrib) done lemma zcong_cancel: @@ -188,7 +188,7 @@ lemma zcong_cancel2: "0 \ m ==> zgcd k m = 1 ==> [k * a = k * b] (mod m) = [a = b] (mod m)" - by (simp add: mult_commute zcong_cancel) + by (simp add: mult.commute zcong_cancel) lemma zcong_zgcd_zmult_zmod: "[a = b] (mod m) ==> [a = b] (mod n) ==> zgcd m n = 1 @@ -197,9 +197,9 @@ apply (subgoal_tac "m dvd n * ka") apply (subgoal_tac "m dvd ka") apply (case_tac [2] "0 \ ka") - apply (metis dvd_mult_div_cancel dvd_refl dvd_mult_left mult_commute zrelprime_zdvd_zmult) - apply (metis abs_dvd_iff abs_of_nonneg add_0 zgcd_0_left zgcd_commute zgcd_zadd_zmult zgcd_zdvd_zgcd_zmult zgcd_zmult_distrib2_abs mult_1_right mult_commute) - apply (metis mult_le_0_iff zdvd_mono zdvd_mult_cancel dvd_triv_left zero_le_mult_iff order_antisym linorder_linear order_refl mult_commute zrelprime_zdvd_zmult) + apply (metis dvd_mult_div_cancel dvd_refl dvd_mult_left mult.commute zrelprime_zdvd_zmult) + apply (metis abs_dvd_iff abs_of_nonneg add_0 zgcd_0_left zgcd_commute zgcd_zadd_zmult zgcd_zdvd_zgcd_zmult zgcd_zmult_distrib2_abs mult_1_right mult.commute) + apply (metis mult_le_0_iff zdvd_mono zdvd_mult_cancel dvd_triv_left zero_le_mult_iff order_antisym linorder_linear order_refl mult.commute zrelprime_zdvd_zmult) apply (metis dvd_triv_left) done @@ -208,7 +208,7 @@ a < m ==> 0 \ b ==> b < m ==> [a = b] (mod m) ==> a = b" apply (unfold zcong_def dvd_def, auto) apply (drule_tac f = "\z. z mod m" in arg_cong) - apply (metis diff_add_cancel mod_pos_pos_trivial add_0 add_commute zmod_eq_0_iff mod_add_right_eq) + apply (metis diff_add_cancel mod_pos_pos_trivial add_0 add.commute zmod_eq_0_iff mod_add_right_eq) done lemma zcong_square_zless: @@ -261,7 +261,7 @@ apply (rule_tac m = m in zcong_zmod_aux) apply (rule trans) apply (rule_tac [2] k = m and m = "a div m - b div m" in zdvd_reduce) - apply (simp add: add_commute) + apply (simp add: add.commute) done lemma zcong_zmod_eq: "0 < m ==> [a = b] (mod m) = (a mod m = b mod m)" @@ -341,14 +341,14 @@ lemma xzgcda_linear_aux1: "(a - r * b) * m + (c - r * d) * (n::int) = (a * m + c * n) - r * (b * m + d * n)" - by (simp add: left_diff_distrib distrib_left mult_assoc) + by (simp add: left_diff_distrib distrib_left mult.assoc) lemma xzgcda_linear_aux2: "r' = s' * m + t' * n ==> r = s * m + t * n ==> (r' mod r) = (s' - (r' div r) * s) * m + (t' - (r' div r) * t) * (n::int)" apply (rule trans) apply (rule_tac [2] xzgcda_linear_aux1 [symmetric]) - apply (simp add: eq_diff_eq mult_commute) + apply (simp add: eq_diff_eq mult.commute) done lemma order_le_neq_implies_less: "(x::'a::order) \ y ==> x \ y ==> x < y" @@ -391,7 +391,7 @@ prefer 2 apply simp apply (unfold zcong_def) - apply (simp (no_asm) add: mult_commute) + apply (simp (no_asm) add: mult.commute) done lemma zcong_lineq_unique: @@ -407,7 +407,7 @@ apply (cut_tac a = a and n = n in zcong_lineq_ex, auto) apply (rule_tac x = "x * b mod n" in exI, safe) apply (simp_all (no_asm_simp)) - apply (metis zcong_scalar zcong_zmod mod_mult_right_eq mult_1 mult_assoc) + apply (metis zcong_scalar zcong_zmod mod_mult_right_eq mult_1 mult.assoc) done end diff -r de51a86fc903 -r cc97b347b301 src/HOL/Old_Number_Theory/Legacy_GCD.thy --- a/src/HOL/Old_Number_Theory/Legacy_GCD.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Old_Number_Theory/Legacy_GCD.thy Fri Jul 04 20:18:47 2014 +0200 @@ -161,7 +161,7 @@ apply (rule_tac n = k in relprime_dvd_mult) apply (simp add: gcd_assoc) apply (simp add: gcd_commute) - apply (simp_all add: mult_commute) + apply (simp_all add: mult.commute) done @@ -173,19 +173,19 @@ lemma gcd_add2 [simp, algebra]: "gcd m (m + n) = gcd m n" proof - have "gcd m (m + n) = gcd (m + n) m" by (rule gcd_commute) - also have "... = gcd (n + m) m" by (simp add: add_commute) + also have "... = gcd (n + m) m" by (simp add: add.commute) also have "... = gcd n m" by simp also have "... = gcd m n" by (rule gcd_commute) finally show ?thesis . qed lemma gcd_add2' [simp, algebra]: "gcd m (n + m) = gcd m n" - apply (subst add_commute) + apply (subst add.commute) apply (rule gcd_add2) done lemma gcd_add_mult[algebra]: "gcd m (k * m + n) = gcd m n" - by (induct k) (simp_all add: add_assoc) + by (induct k) (simp_all add: add.assoc) lemma gcd_dvd_prod: "gcd m n dvd m * n" using mult_dvd_mono [of 1] by auto @@ -351,7 +351,7 @@ hence "a * ((b - 1) * y) = d + (x*b*(b - 1) - d*b)" by (simp only: diff_add_assoc[OF dble, of d, symmetric]) hence "a * ((b - 1) * y) = b*(x*(b - 1) - d) + d" - by (simp only: diff_mult_distrib2 add_commute mult_ac) + by (simp only: diff_mult_distrib2 add.commute mult_ac) hence ?thesis using H(1,2) apply - apply (rule exI[where x=d], simp) @@ -374,7 +374,7 @@ hence "a * x * k - b * y*k = d*k \ b * x * k - a * y*k = d*k" by (algebra add: diff_mult_distrib) hence "a * (x * k) - b * (y*k) = ?g \ b * (x * k) - a * (y*k) = ?g" - by (simp add: k mult_assoc) + by (simp add: k mult.assoc) thus ?thesis by blast qed @@ -392,7 +392,7 @@ qed lemma gcd_mult_distrib: "gcd(a * c) (b * c) = c * gcd a b" -by(simp add: gcd_mult_distrib2 mult_commute) +by(simp add: gcd_mult_distrib2 mult.commute) lemma gcd_bezout: "(\x y. a * x - b * y = d \ b * x - a * y = d) \ gcd a b dvd d" (is "?lhs \ ?rhs") @@ -405,7 +405,7 @@ hence "a * x*k - b * y*k = ?g*k \ b * x * k - a * y*k = ?g*k" by (simp only: diff_mult_distrib) hence "a * (x*k) - b * (y*k) = d \ b * (x * k) - a * (y*k) = d" - by (simp add: k[symmetric] mult_assoc) + by (simp add: k[symmetric] mult.assoc) hence ?lhs by blast} moreover {fix x y assume H: "a * x - b * y = d \ b * x - a * y = d" @@ -426,7 +426,7 @@ qed lemma gcd_mult': "gcd b (a * b) = b" -by (simp add: mult_commute[of a b]) +by (simp add: mult.commute[of a b]) lemma gcd_add: "gcd(a + b) b = gcd a b" "gcd(b + a) b = gcd a b" "gcd a (a + b) = gcd a b" "gcd a (b + a) = gcd a b" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Old_Number_Theory/Pocklington.thy --- a/src/HOL/Old_Number_Theory/Pocklington.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Old_Number_Theory/Pocklington.thy Fri Jul 04 20:18:47 2014 +0200 @@ -120,7 +120,7 @@ lemma cong_mult_rcancel: assumes an: "coprime a n" and axy:"[x*a = y*a] (mod n)" shows "[x = y] (mod n)" - using cong_mult_lcancel[OF an axy[unfolded mult_commute[of _a]]] . + using cong_mult_lcancel[OF an axy[unfolded mult.commute[of _a]]] . lemma cong_refl: "[x = x] (mod n)" by (simp add: modeq_def) @@ -177,13 +177,13 @@ proof assume H: "?rhs" from cong_mult[OF cong_refl[of a n] H] show ?lhs . next - assume H: "?lhs" hence H': "[x*a = y*a] (mod n)" by (simp add: mult_commute) + assume H: "?lhs" hence H': "[x*a = y*a] (mod n)" by (simp add: mult.commute) from cong_mult_rcancel[OF an H'] show ?rhs . qed lemma cong_mult_rcancel_eq: assumes an: "coprime a n" shows "[x * a = y * a] (mod n) \ [x = y] (mod n)" -using cong_mult_lcancel_eq[OF an, of x y] by (simp add: mult_commute) +using cong_mult_lcancel_eq[OF an, of x y] by (simp add: mult.commute) lemma cong_add_lcancel_eq: "[a + x = a + y] (mod n) \ [x = y] (mod n)" by (simp add: nat_mod) @@ -358,7 +358,7 @@ also have "\ = m mod a" by (simp add: mod_mult2_eq) finally have th1: "[?w = m] (mod a)" by (simp add: modeq_def) from xq12(2) have "?w mod b = ((n + q2 * b) mod (a*b)) mod b" by simp - also have "\ = ((n + q2 * b) mod (b*a)) mod b" by (simp add: mult_commute) + also have "\ = ((n + q2 * b) mod (b*a)) mod b" by (simp add: mult.commute) also have "\ = n mod b" by (simp add: mod_mult2_eq) finally have th2: "[?w = n] (mod b)" by (simp add: modeq_def) {fix y assume H: "y < a*b" "[y = m] (mod a)" "[y = n] (mod b)" @@ -638,7 +638,7 @@ also have "[(\i\?S. ?h i) = ?P] (mod n)" using eq0 fS an by (simp add: setprod_def modeq_def) finally show "[?P*a^ (\ n) = ?P*1] (mod n)" - unfolding cardfS mult_commute[of ?P "a^ (card ?S)"] + unfolding cardfS mult.commute[of ?P "a^ (card ?S)"] nproduct_cmul[OF fS, symmetric] mult_1_right by simp qed from cong_mult_lcancel[OF nP Paphi] have ?thesis . } @@ -856,7 +856,7 @@ from H have onz: "?o \ 0" by (simp add: ord_eq_0) hence op: "?o > 0" by simp from mod_div_equality[of d "ord n a"] lh - have "[a^(?o*?q + ?r) = 1] (mod n)" by (simp add: modeq_def mult_commute) + have "[a^(?o*?q + ?r) = 1] (mod n)" by (simp add: modeq_def mult.commute) hence "[(a^?o)^?q * (a^?r) = 1] (mod n)" by (simp add: modeq_def power_mult[symmetric] power_add[symmetric]) hence th: "[a^?r = 1] (mod n)" @@ -964,7 +964,7 @@ from H[rule_format, of d] h d have "d = 1" by blast} moreover {assume h: "e\<^sup>2 \ n" - from e have "e dvd n" unfolding dvd_def by (simp add: mult_commute) + from e have "e dvd n" unfolding dvd_def by (simp add: mult.commute) with H[rule_format, of e] h have "e=1" by simp with e have "d = n" by simp} ultimately have "d=1 \ d=n" by blast} @@ -1231,7 +1231,7 @@ from prime_ge_2[OF p(1)] have p01: "p \ 0" "p \ 1" "p =Suc(p - 1)" by arith+ from div_mult1_eq[of r q p] p(2) have eq1: "r* (q div p) = (n - 1) div p" - unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult_commute) + unfolding qrn[symmetric] dvd_eq_mod_eq_0 by (simp add: mult.commute) have ath: "\a (b::nat). a <= b \ a \ 0 ==> 1 <= a \ 1 <= b" by arith from n0 have n00: "n \ 0" by arith from mod_le[OF n00] @@ -1246,7 +1246,7 @@ with eq0 have "a^ (n - 1) = (n*s)^p" by (simp add: power_mult[symmetric]) hence "1 = (n*s)^(Suc (p - 1)) mod n" using bqn p01 by simp - also have "\ = 0" by (simp add: mult_assoc) + also have "\ = 0" by (simp add: mult.assoc) finally have False by simp } then have th11: "a ^ ((n - 1) div p) mod n \ 0" by auto have th1: "[a ^ ((n - 1) div p) mod n = a ^ ((n - 1) div p)] (mod n)" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Old_Number_Theory/Primes.thy --- a/src/HOL/Old_Number_Theory/Primes.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Old_Number_Theory/Primes.thy Fri Jul 04 20:18:47 2014 +0200 @@ -108,7 +108,7 @@ declare nat_mult_dvd_cancel_disj[presburger] lemma nat_mult_dvd_cancel_disj'[presburger]: - "(m\nat)*k dvd n*k \ k = 0 \ m dvd n" unfolding mult_commute[of m k] mult_commute[of n k] by presburger + "(m\nat)*k dvd n*k \ k = 0 \ m dvd n" unfolding mult.commute[of m k] mult.commute[of n k] by presburger lemma divides_mul_l: "(a::nat) dvd b ==> (c * a) dvd (c * b)" by presburger @@ -120,7 +120,7 @@ lemma divides_div_not: "(x::nat) = (q * n) + r \ 0 < r \ r < n ==> ~(n dvd x)" proof(auto simp add: dvd_def) fix k assume H: "0 < r" "r < n" "q * n + r = n * k" - from H(3) have r: "r = n* (k -q)" by(simp add: diff_mult_distrib2 mult_commute) + from H(3) have r: "r = n* (k -q)" by(simp add: diff_mult_distrib2 mult.commute) {assume "k - q = 0" with r H(1) have False by simp} moreover {assume "k - q \ 0" with r have "r \ n" by auto @@ -186,7 +186,7 @@ using coprime_def gcd_bezout by auto lemma coprime_divprod: "d dvd a * b \ coprime d a \ d dvd b" - using relprime_dvd_mult_iff[of d a b] by (auto simp add: coprime_def mult_commute) + using relprime_dvd_mult_iff[of d a b] by (auto simp add: coprime_def mult.commute) lemma coprime_1[simp]: "coprime a 1" by (simp add: coprime_def) lemma coprime_1'[simp]: "coprime 1 a" by (simp add: coprime_def) @@ -205,7 +205,7 @@ from bezout_gcd_strong[OF az, of b] obtain x y where xy: "a*x = b*y + ?g" by blast from xy a b have "?g * a'*x = ?g * (b'*y + 1)" by (simp add: algebra_simps) - hence "?g * (a'*x) = ?g * (b'*y + 1)" by (simp add: mult_assoc) + hence "?g * (a'*x) = ?g * (b'*y + 1)" by (simp add: mult.assoc) hence "a'*x = (b'*y + 1)" by (simp only: nat_mult_eq_cancel1[OF z']) hence "a'*x - b'*y = 1" by simp @@ -296,7 +296,7 @@ using z by auto then have "a^n * x - b^n * y = ?g^n \ b^n * x - a^n * y = ?g^n" using z ab'' by (simp only: power_mult_distrib[symmetric] - diff_mult_distrib2 mult_assoc[symmetric]) + diff_mult_distrib2 mult.assoc[symmetric]) hence ?thesis by blast } ultimately show ?thesis by blast qed @@ -334,7 +334,7 @@ from ab'(1) have "a' dvd a" unfolding dvd_def by blast with dc have th0: "a' dvd b*c" using dvd_trans[of a' a "b*c"] by simp from dc ab'(1,2) have "a'*?g dvd (b'*?g) *c" by auto - hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult_assoc) + hence "?g*a' dvd ?g * (b' * c)" by (simp add: mult.assoc) with z have th_1: "a' dvd b'*c" by simp from coprime_divprod[OF th_1 ab'(3)] have thc: "a' dvd c" . from ab' have "a = ?g*a'" by algebra @@ -355,11 +355,11 @@ from gcd_coprime_exists[OF z] obtain a' b' where ab': "a = a' * ?g" "b = b' * ?g" "coprime a' b'" by blast from ab have "(a' * ?g) ^ n dvd (b' * ?g)^n" by (simp add: ab'(1,2)[symmetric]) - hence "?g^n*a'^n dvd ?g^n *b'^n" by (simp only: power_mult_distrib mult_commute) + hence "?g^n*a'^n dvd ?g^n *b'^n" by (simp only: power_mult_distrib mult.commute) with zn z n have th0:"a'^n dvd b'^n" by (auto simp add: nat_power_eq_0_iff) have "a' dvd a'^n" by (simp add: m) with th0 have "a' dvd b'^n" using dvd_trans[of a' "a'^n" "b'^n"] by simp - hence th1: "a' dvd b'^m * b'" by (simp add: m mult_commute) + hence th1: "a' dvd b'^m * b'" by (simp add: m mult.commute) from coprime_divprod[OF th1 coprime_exp[OF ab'(3), of m]] have "a' dvd b'" . hence "a'*?g dvd b'*?g" by simp @@ -372,7 +372,7 @@ proof- from mr nr obtain m' n' where m': "r = m*m'" and n': "r = n*n'" unfolding dvd_def by blast - from mr n' have "m dvd n'*n" by (simp add: mult_commute) + from mr n' have "m dvd n'*n" by (simp add: mult.commute) hence "m dvd n'" using relprime_dvd_mult_iff[OF mn[unfolded coprime_def]] by simp then obtain k where k: "n' = m*k" unfolding dvd_def by blast from n' k show ?thesis unfolding dvd_def by auto @@ -609,7 +609,7 @@ have "p dvd a \ p dvd b" . moreover {assume pa: "p dvd a" - have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute) + have pnba: "p^n dvd b*a" using pab by (simp add: mult.commute) from coprime_prime[OF ab, of p] p pa have "\ p dvd b" by blast with prime_coprime[OF p, of b] b have cpb: "coprime b p" using coprime_commute by blast @@ -618,7 +618,7 @@ from coprime_divprod[OF pnba pnb] have ?thesis by blast } moreover {assume pb: "p dvd b" - have pnba: "p^n dvd b*a" using pab by (simp add: mult_commute) + have pnba: "p^n dvd b*a" using pab by (simp add: mult.commute) from coprime_prime[OF ab, of p] p pb have "\ p dvd a" by blast with prime_coprime[OF p, of a] a have cpb: "coprime a p" using coprime_commute by blast @@ -632,7 +632,7 @@ lemma nat_mult_eq_one: "(n::nat) * m = 1 \ n = 1 \ m = 1" (is "?lhs \ ?rhs") proof assume H: "?lhs" - hence "n dvd 1" "m dvd 1" unfolding dvd_def by (auto simp add: mult_commute) + hence "n dvd 1" "m dvd 1" unfolding dvd_def by (auto simp add: mult.commute) thus ?rhs by auto next assume ?rhs then show ?lhs by auto @@ -705,7 +705,7 @@ from coprime_lmul2 [OF H(2)[unfolded k coprime_commute[of "p^n*k" a]]] have kb: "coprime k a" by (simp add: coprime_commute) from H(3) l k pn0 n have kbln: "k * a = l ^ n" - by (simp add: power_mult_distrib mult_commute) + by (simp add: power_mult_distrib mult.commute) from H(1)[rule_format, OF lc kb kbln] obtain r s where rs: "k = r ^n" "a = s^n" by blast from k rs(1) have "b = (p*r)^n" by (simp add: power_mult_distrib) @@ -772,7 +772,7 @@ moreover {assume px: "p dvd y" then obtain d where d: "y = p*d" unfolding dvd_def by blast - from Suc.prems d have "p*d*x = p^Suc k" by (simp add: mult_commute) + from Suc.prems d have "p*d*x = p^Suc k" by (simp add: mult.commute) hence th: "d*x = p^k" using p0 by simp from Suc.hyps[OF th] obtain i j where ij: "d = p^i" "x = p^j" by blast with d have "y = p^Suc i" by simp @@ -800,7 +800,7 @@ shows "d dvd p^k \ (\ i. i \ k \ d = p ^i)" proof assume H: "d dvd p^k" then obtain e where e: "d*e = p^k" - unfolding dvd_def apply (auto simp add: mult_commute) by blast + unfolding dvd_def apply (auto simp add: mult.commute) by blast from prime_power_mult[OF p e] obtain i j where ij: "d = p^i" "e=p^j" by blast from prime_ge_2[OF p] have p1: "p > 1" by arith from e ij have "p^(i + j) = p^k" by (simp add: power_add) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Old_Number_Theory/Quadratic_Reciprocity.thy --- a/src/HOL/Old_Number_Theory/Quadratic_Reciprocity.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Old_Number_Theory/Quadratic_Reciprocity.thy Fri Jul 04 20:18:47 2014 +0200 @@ -22,7 +22,7 @@ from finite_A have "a * setsum id A = setsum (%x. a * x) A" by (auto simp add: setsum_const_mult id_def) also have "setsum (%x. a * x) = setsum (%x. x * a)" - by (auto simp add: mult_commute) + by (auto simp add: mult.commute) also have "setsum (%x. x * a) A = setsum id B" by (simp add: B_def setsum.reindex [OF inj_on_xa_A]) also have "... = setsum (%x. p * (x div p) + StandardRes p x) B" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Old_Number_Theory/WilsonBij.thy --- a/src/HOL/Old_Number_Theory/WilsonBij.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Old_Number_Theory/WilsonBij.thy Fri Jul 04 20:18:47 2014 +0200 @@ -207,7 +207,7 @@ lemma reciP_sym: "zprime p ==> symP (reciR p)" apply (unfold reciR_def symP_def) - apply (simp add: mult_commute) + apply (simp add: mult.commute) apply auto done @@ -234,7 +234,7 @@ apply (subst setprod.insert) apply (auto simp add: fin_bijER) apply (subgoal_tac "zcong ((a * b) * \A) (1 * 1) p") - apply (simp add: mult_assoc) + apply (simp add: mult.assoc) apply (rule zcong_zmult) apply auto done diff -r de51a86fc903 -r cc97b347b301 src/HOL/Old_Number_Theory/WilsonRuss.thy --- a/src/HOL/Old_Number_Theory/WilsonRuss.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Old_Number_Theory/WilsonRuss.thy Fri Jul 04 20:18:47 2014 +0200 @@ -256,7 +256,7 @@ apply (subgoal_tac [5] "zcong (a * inv p a * (\x\wset (a - 1) p. x)) (1 * 1) p") prefer 5 - apply (simp add: mult_assoc) + apply (simp add: mult.assoc) apply (rule_tac [5] zcong_zmult) apply (rule_tac [5] inv_is_inv) apply (tactic "clarify_tac @{context} 4") diff -r de51a86fc903 -r cc97b347b301 src/HOL/Power.thy --- a/src/HOL/Power.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Power.thy Fri Jul 04 20:18:47 2014 +0200 @@ -53,7 +53,7 @@ lemma power_commutes: "a ^ n * a = a * a ^ n" - by (induct n) (simp_all add: mult_assoc) + by (induct n) (simp_all add: mult.assoc) lemma power_Suc2: "a ^ Suc n = a ^ n * a" @@ -71,11 +71,11 @@ by (simp add: numeral_2_eq_2) lemma power3_eq_cube: "a ^ 3 = a * a * a" - by (simp add: numeral_3_eq_3 mult_assoc) + by (simp add: numeral_3_eq_3 mult.assoc) lemma power_even_eq: "a ^ (2 * n) = (a ^ n)\<^sup>2" - by (subst mult_commute) (simp add: power_mult) + by (subst mult.commute) (simp add: power_mult) lemma power_odd_eq: "a ^ Suc (2*n) = a * (a ^ n)\<^sup>2" @@ -88,7 +88,7 @@ lemma power_numeral_odd: "z ^ numeral (Num.Bit1 w) = (let w = z ^ (numeral w) in z * w * w)" unfolding numeral_Bit1 One_nat_def add_Suc_right add_0_right - unfolding power_Suc power_add Let_def mult_assoc .. + unfolding power_Suc power_add Let_def mult.assoc .. lemma funpow_times_power: "(times x ^^ f x) = times (x ^ f x)" @@ -100,7 +100,7 @@ with Suc have "n = g x" by simp with Suc have "times x ^^ g x = times (x ^ g x)" by simp moreover from Suc g_def have "f x = g x + 1" by simp - ultimately show ?case by (simp add: power_add funpow_add fun_eq_iff mult_assoc) + ultimately show ?case by (simp add: power_add funpow_add fun_eq_iff mult.assoc) qed end @@ -197,7 +197,7 @@ case 0 show ?case by simp next case (Suc n) then show ?case - by (simp del: power_Suc add: power_Suc2 mult_assoc) + by (simp del: power_Suc add: power_Suc2 mult.assoc) qed lemma power_minus_Bit0: @@ -626,7 +626,7 @@ lemma power2_diff: fixes x y :: "'a::comm_ring_1" shows "(x - y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 - 2 * x * y" - by (simp add: ring_distribs power2_eq_square mult_2) (rule mult_commute) + by (simp add: ring_distribs power2_eq_square mult_2) (rule mult.commute) lemma power_0_Suc [simp]: "(0::'a::{power, semiring_0}) ^ Suc n = 0" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Presburger.thy --- a/src/HOL/Presburger.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Presburger.thy Fri Jul 04 20:18:47 2014 +0200 @@ -54,8 +54,8 @@ "(d::'a::{comm_ring,Rings.dvd}) dvd D \ \x k. (\d dvd x + t) = (\d dvd (x - k*D) + t)" "\x k. F = F" apply (auto elim!: dvdE simp add: algebra_simps) -unfolding mult_assoc [symmetric] distrib_right [symmetric] left_diff_distrib [symmetric] -unfolding dvd_def mult_commute [of d] +unfolding mult.assoc [symmetric] distrib_right [symmetric] left_diff_distrib [symmetric] +unfolding dvd_def mult.commute [of d] by auto subsection{* The A and B sets *} diff -r de51a86fc903 -r cc97b347b301 src/HOL/Probability/Bochner_Integration.thy --- a/src/HOL/Probability/Bochner_Integration.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Probability/Bochner_Integration.thy Fri Jul 04 20:18:47 2014 +0200 @@ -743,7 +743,7 @@ finally have s_fin: "(\\<^sup>+x. norm (s i x) \M) < \" . have "(\\<^sup>+ x. norm (f x) \M) \ (\\<^sup>+ x. ereal (norm (f x - s i x)) + ereal (norm (s i x)) \M)" - by (auto intro!: nn_integral_mono) (metis add_commute norm_triangle_sub) + by (auto intro!: nn_integral_mono) (metis add.commute norm_triangle_sub) also have "\ = (\\<^sup>+x. norm (f x - s i x) \M) + (\\<^sup>+x. norm (s i x) \M)" by (rule nn_integral_add) auto also have "\ < \" @@ -779,7 +779,7 @@ (auto intro: s simple_bochner_integrable_compose2) also have "\ \ (\\<^sup>+x. ereal (norm (f x - s n x)) + norm (f x) \M)" by (auto intro!: nn_integral_mono) - (metis add_commute norm_minus_commute norm_triangle_sub) + (metis add.commute norm_minus_commute norm_triangle_sub) also have "\ = ?t n" by (rule nn_integral_add) auto finally show "norm (?s n) \ ?t n" . diff -r de51a86fc903 -r cc97b347b301 src/HOL/Probability/Convolution.thy --- a/src/HOL/Probability/Convolution.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Probability/Convolution.thy Fri Jul 04 20:18:47 2014 +0200 @@ -66,7 +66,7 @@ apply (subst nn_integral_indicator[symmetric], simp) apply (subst nn_integral_convolution, auto intro!: borel_measurable_indicator borel_measurable_indicator' convolution_finite)+ - by (rule nn_integral_cong)+ (auto simp: semigroup_add_class.add_assoc) + by (rule nn_integral_cong)+ (auto simp: semigroup_add_class.add.assoc) lemma convolution_emeasure_3': assumes [simp, measurable]:"A \ sets borel" @@ -96,7 +96,7 @@ have "emeasure (M \ N) A = \\<^sup>+x. \\<^sup>+y. indicator A (x + y) \N \M" by (auto intro!: convolution_emeasure') also have "... = \\<^sup>+x. \\<^sup>+y. (\(x,y). indicator A (x + y)) (x, y) \N \M" by (auto intro!: nn_integral_cong) also have "... = \\<^sup>+y. \\<^sup>+x. (\(x,y). indicator A (x + y)) (x, y) \M \N" by (rule Fubini[symmetric]) simp - also have "... = emeasure (N \ M) A" by (auto intro!: nn_integral_cong simp: add_commute convolution_emeasure') + also have "... = emeasure (N \ M) A" by (auto intro!: nn_integral_cong simp: add.commute convolution_emeasure') finally show "emeasure (M \ N) A = emeasure (N \ M) A" by simp qed diff -r de51a86fc903 -r cc97b347b301 src/HOL/Probability/Distributions.thy --- a/src/HOL/Probability/Distributions.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Probability/Distributions.thy Fri Jul 04 20:18:47 2014 +0200 @@ -865,7 +865,7 @@ proof (subst nn_integral_FTC_atLeast) have "((\a. - (exp (- (a\<^sup>2 * (1 + s\<^sup>2))) / (2 + 2 * s\<^sup>2))) ---> (- (0 / (2 + 2 * s\<^sup>2)))) at_top" apply (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_compose[OF filterlim_uminus_at_bot_at_top]) - apply (subst mult_commute) + apply (subst mult.commute) apply (auto intro!: filterlim_tendsto_pos_mult_at_top filterlim_at_top_mult_at_top[OF filterlim_ident filterlim_ident] simp: add_pos_nonneg power2_eq_square add_nonneg_eq_0_iff) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Probability/Finite_Product_Measure.thy --- a/src/HOL/Probability/Finite_Product_Measure.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Probability/Finite_Product_Measure.thy Fri Jul 04 20:18:47 2014 +0200 @@ -643,7 +643,7 @@ using E by (subst insert) (auto intro!: setprod.cong) also have "(\j\I. if j \ J - {i} then emeasure (M j) (E j) else emeasure (M j) (space (M j))) * emeasure (M i) (if i \ J then E i else space (M i)) = (\j\insert i I. ?f J E j)" - using insert by (auto simp: mult_commute intro!: arg_cong2[where f="op *"] setprod.cong) + using insert by (auto simp: mult.commute intro!: arg_cong2[where f="op *"] setprod.cong) also have "\ = (\j\J \ ?I. ?f J E j)" using insert(1,2) J E by (intro setprod.mono_neutral_right) auto finally show "?\ ?p = \" . diff -r de51a86fc903 -r cc97b347b301 src/HOL/Probability/Independent_Family.thy --- a/src/HOL/Probability/Independent_Family.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Probability/Independent_Family.thy Fri Jul 04 20:18:47 2014 +0200 @@ -1301,7 +1301,7 @@ unfolding indep_var_def proof - have *: "(\\. X1 \ * X2 \) = (\\. \i\UNIV. (case_bool X1 X2 i \))" - by (simp add: UNIV_bool mult_commute) + by (simp add: UNIV_bool mult.commute) have **: "(\ _. borel) = case_bool borel borel" by (rule ext, metis (full_types) bool.simps(3) bool.simps(4)) show ?eq @@ -1310,7 +1310,7 @@ apply (auto) apply (subst **, subst indep_var_def [symmetric], rule assms) apply (simp split: bool.split add: assms) - by (simp add: UNIV_bool mult_commute) + by (simp add: UNIV_bool mult.commute) show ?int apply (subst *) apply (rule indep_vars_integrable) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Probability/Infinite_Product_Measure.thy --- a/src/HOL/Probability/Infinite_Product_Measure.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Probability/Infinite_Product_Measure.thy Fri Jul 04 20:18:47 2014 +0200 @@ -565,7 +565,7 @@ also have "emeasure S ?E = (\j\J \ {..j\J \ {..j\J - (J \ {..j\J. emeasure M (E j))" - by (subst mult_commute) (auto simp: J setprod.subset_diff[symmetric]) + by (subst mult.commute) (auto simp: J setprod.subset_diff[symmetric]) finally show "emeasure ?D ?X = (\j\J. emeasure M (E j))" . qed simp_all diff -r de51a86fc903 -r cc97b347b301 src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Fri Jul 04 20:18:47 2014 +0200 @@ -1080,7 +1080,7 @@ by (auto intro: continuous_at_imp_continuous_on) qed simp then show ?thesis - by (auto simp: mult_commute) + by (auto simp: mult.commute) qed text {* @@ -1160,7 +1160,7 @@ using cont by (intro continuous_at_imp_continuous_on) auto show ?has ?eq using has_bochner_integral_FTC_Icc[OF `a \ b` 1 2] integral_FTC_Icc[OF `a \ b` 1 2] - by (auto simp: mult_commute) + by (auto simp: mult.commute) qed lemma nn_integral_FTC_atLeast: diff -r de51a86fc903 -r cc97b347b301 src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy --- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Fri Jul 04 20:18:47 2014 +0200 @@ -614,7 +614,7 @@ using f by (intro simple_function_partition) auto also have "\ = c * integral\<^sup>S M f" using f unfolding simple_integral_def - by (subst setsum_ereal_right_distrib) (auto simp: emeasure_nonneg mult_assoc Int_def conj_commute) + by (subst setsum_ereal_right_distrib) (auto simp: emeasure_nonneg mult.assoc Int_def conj_commute) finally show ?thesis . qed @@ -1129,7 +1129,7 @@ lemma nn_integral_multc: assumes "f \ borel_measurable M" "0 \ c" shows "(\\<^sup>+ x. f x * c \M) = integral\<^sup>N M f * c" - unfolding mult_commute[of _ c] nn_integral_cmult[OF assms] by simp + unfolding mult.commute[of _ c] nn_integral_cmult[OF assms] by simp lemma nn_integral_indicator[simp]: "A \ sets M \ (\\<^sup>+ x. indicator A x\M) = (emeasure M) A" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Probability/Sigma_Algebra.thy Fri Jul 04 20:18:47 2014 +0200 @@ -830,11 +830,11 @@ case 0 show ?case by simp next case (Suc m) thus ?case - by (metis Suc_eq_plus1 assms nat_add_commute nat_add_left_commute subset_trans) + by (metis Suc_eq_plus1 assms add.commute add.left_commute subset_trans) qed } hence "!!m n. m < n \ A m \ A n" - by (metis add_commute le_add_diff_inverse nat_less_le) + by (metis add.commute le_add_diff_inverse nat_less_le) thus ?thesis by (auto simp add: disjoint_family_on_def) (metis insert_absorb insert_subset le_SucE le_antisym not_leE) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Proofs/Extraction/Euclid.thy --- a/src/HOL/Proofs/Extraction/Euclid.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Proofs/Extraction/Euclid.thy Fri Jul 04 20:18:47 2014 +0200 @@ -114,12 +114,12 @@ assume "m \ n" with `0 < m` have "m dvd fact n" by (rule Suc) then have "m dvd (fact n * Suc n)" by (rule dvd_mult2) - then show ?thesis by (simp add: mult_commute) + then show ?thesis by (simp add: mult.commute) next assume "m = Suc n" then have "m dvd (fact n * Suc n)" by (auto intro: dvdI simp: mult_ac) - then show ?thesis by (simp add: mult_commute) + then show ?thesis by (simp add: mult.commute) qed qed diff -r de51a86fc903 -r cc97b347b301 src/HOL/Proofs/Extraction/Greatest_Common_Divisor.thy --- a/src/HOL/Proofs/Extraction/Greatest_Common_Divisor.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Proofs/Extraction/Greatest_Common_Divisor.thy Fri Jul 04 20:18:47 2014 +0200 @@ -39,7 +39,7 @@ by iprover have mn: "Suc m < n" by (rule 1) from h1 h1' h2' Suc have "k * (m1 * q + r1) = n" - by (simp add: add_mult_distrib2 nat_mult_assoc [symmetric]) + by (simp add: add_mult_distrib2 mult.assoc [symmetric]) moreover have "\l l1 l2. l * l1 = n \ l * l2 = Suc m \ l \ k" proof - fix l l1 l2 diff -r de51a86fc903 -r cc97b347b301 src/HOL/Proofs/Lambda/ListApplication.thy --- a/src/HOL/Proofs/Lambda/ListApplication.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Proofs/Lambda/ListApplication.thy Fri Jul 04 20:18:47 2014 +0200 @@ -110,7 +110,7 @@ prefer 2 apply (erule allE, erule mp, rule refl) apply simp - apply (simp only: foldl_conv_fold add_commute fold_plus_listsum_rev) + apply (simp only: foldl_conv_fold add.commute fold_plus_listsum_rev) apply (fastforce simp add: listsum_map_remove1) apply clarify apply (rule assms) @@ -126,7 +126,7 @@ apply (rule le_imp_less_Suc) apply (rule trans_le_add1) apply (rule trans_le_add2) - apply (simp only: foldl_conv_fold add_commute fold_plus_listsum_rev) + apply (simp only: foldl_conv_fold add.commute fold_plus_listsum_rev) apply (simp add: member_le_listsum_nat) done diff -r de51a86fc903 -r cc97b347b301 src/HOL/Quotient_Examples/Int_Pow.thy --- a/src/HOL/Quotient_Examples/Int_Pow.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Quotient_Examples/Int_Pow.thy Fri Jul 04 20:18:47 2014 +0200 @@ -53,7 +53,7 @@ proof (cases "m\n") have [simp]: "a \ carrier G" using `a \ _` by (rule Units_closed) case True - then obtain k where *:"m = k + n" and **:"m = n + k" by (metis Nat.le_iff_add nat_add_commute) + then obtain k where *:"m = k + n" and **:"m = n + k" by (metis Nat.le_iff_add add.commute) have "a (^) (m::nat) \ inv (a (^) (n::nat)) = a (^) k" using * by (auto simp add: nat_pow_mult[symmetric] m_assoc) also have "\ = inv (a (^) n) \ a (^) m" @@ -63,7 +63,7 @@ have [simp]: "a \ carrier G" using `a \ _` by (rule Units_closed) case False then obtain k where *:"n = k + m" and **:"n = m + k" - by (metis Nat.le_iff_add nat_add_commute nat_le_linear) + by (metis Nat.le_iff_add add.commute nat_le_linear) have "a (^) (m::nat) \ inv (a (^) (n::nat)) = inv(a (^) k)" using * by (auto simp add: nat_pow_mult[symmetric] m_assoc[symmetric] inv_mult_units) also have "\ = inv (a (^) n) \ a (^) m" @@ -86,7 +86,7 @@ proof(cases "b\c") have [simp]: "a \ carrier G" using `a \ _` by (rule Units_closed) case True - then obtain n where "b = n + c" by (metis Nat.le_iff_add nat_add_commute) + then obtain n where "b = n + c" by (metis Nat.le_iff_add add.commute) then have "d = n + e" using eq by arith from `b = _` have "a (^) b \ inv (a (^) c) = a (^) n" by (auto simp add: nat_pow_mult[symmetric] m_assoc) @@ -96,7 +96,7 @@ next have [simp]: "a \ carrier G" using `a \ _` by (rule Units_closed) case False - then obtain n where "c = n + b" by (metis Nat.le_iff_add nat_add_commute nat_le_linear) + then obtain n where "c = n + b" by (metis Nat.le_iff_add add.commute nat_le_linear) then have "e = n + d" using eq by arith from `c = _` have "a (^) b \ inv (a (^) c) = inv (a (^) n)" by (auto simp add: nat_pow_mult[symmetric] m_assoc[symmetric] inv_mult_units) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Quotient_Examples/Quotient_Rat.thy --- a/src/HOL/Quotient_Examples/Quotient_Rat.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Quotient_Examples/Quotient_Rat.thy Fri Jul 04 20:18:47 2014 +0200 @@ -41,14 +41,14 @@ "times_rat_raw (a :: int, b :: int) (c, d) = (a * c, b * d)" quotient_definition - "(op *) :: (rat \ rat \ rat)" is times_rat_raw by (auto simp add: mult_assoc mult_left_commute) + "(op *) :: (rat \ rat \ rat)" is times_rat_raw by (auto simp add: mult.assoc mult.left_commute) fun plus_rat_raw where "plus_rat_raw (a :: int, b :: int) (c, d) = (a * d + c * b, b * d)" quotient_definition "(op +) :: (rat \ rat \ rat)" is plus_rat_raw - by (auto simp add: mult_commute mult_left_commute int_distrib(2)) + by (auto simp add: mult.commute mult.left_commute int_distrib(2)) fun uminus_rat_raw where "uminus_rat_raw (a :: int, b :: int) = (-a, b)" @@ -78,13 +78,13 @@ have "a * f * b * f * d * d \ e * b * b * f * d * d" using le nz(2) by (metis linorder_le_cases mult_right_mono mult_right_mono_neg) then have "c * f * f * d * (b * b) \ e * f * d * d * (b * b)" using eq - by (metis (no_types) mult_assoc mult_commute) + by (metis (no_types) mult.assoc mult.commute) then have "c * f * f * d \ e * f * d * d" using b2 by (metis leD linorder_le_less_linear mult_strict_right_mono) then have "c * f * f * d * h * h \ e * f * d * d * h * h" using nz(4) by (metis linorder_le_cases mult_right_mono mult_right_mono_neg) then have "c * h * (d * h) * (f * f) \ g * d * (d * h) * (f * f)" using eq - by (metis (no_types) mult_assoc mult_commute) + by (metis (no_types) mult.assoc mult.commute) then have "c * h * (d * h) \ g * d * (d * h)" using f2 by (metis leD linorder_le_less_linear mult_strict_right_mono) } @@ -128,7 +128,7 @@ show "1 * a = a" by partiality_descending auto show "a + b + c = a + (b + c)" - by partiality_descending (auto simp add: mult_commute distrib_left) + by partiality_descending (auto simp add: mult.commute distrib_left) show "a + b = b + a" by partiality_descending auto show "0 + a = a" @@ -138,7 +138,7 @@ show "a - b = a + - b" by (simp add: minus_rat_def) show "(a + b) * c = a * c + b * c" - by partiality_descending (auto simp add: mult_commute distrib_left) + by partiality_descending (auto simp add: mult.commute distrib_left) show "(0 :: rat) \ (1 :: rat)" by partiality_descending auto qed @@ -167,7 +167,7 @@ "rat_inverse_raw (a :: int, b :: int) = (if a = 0 then (0, 1) else (b, a))" quotient_definition - "inverse :: rat \ rat" is rat_inverse_raw by (force simp add: mult_commute) + "inverse :: rat \ rat" is rat_inverse_raw by (force simp add: mult.commute) definition divide_rat_def: "q / r = q * inverse (r::rat)" @@ -194,7 +194,7 @@ { assume "q \ r" and "r \ s" then show "q \ s" - proof (partiality_descending, auto simp add: mult_assoc[symmetric]) + proof (partiality_descending, auto simp add: mult.assoc[symmetric]) fix a b c d e f :: int assume nz: "b \ 0" "d \ 0" "f \ 0" then have d2: "d * d > 0" @@ -220,9 +220,9 @@ show "q \ q" by partiality_descending auto show "(q < r) = (q \ r \ \ r \ q)" unfolding less_rat_def - by partiality_descending (auto simp add: le_less mult_commute) + by partiality_descending (auto simp add: le_less mult.commute) show "q \ r \ r \ q" - by partiality_descending (auto simp add: mult_commute linorder_linear) + by partiality_descending (auto simp add: mult.commute linorder_linear) } qed @@ -232,25 +232,25 @@ proof fix q r s :: rat show "q \ r ==> s + q \ s + r" - proof (partiality_descending, auto simp add: algebra_simps, simp add: mult_assoc[symmetric]) + proof (partiality_descending, auto simp add: algebra_simps, simp add: mult.assoc[symmetric]) fix a b c d e :: int assume "e \ 0" then have e2: "e * e > 0" by (metis linorder_neqE_linordered_idom mult_eq_0_iff not_square_less_zero) assume "a * b * d * d \ b * b * c * d" then show "a * b * d * d * e * e * e * e \ b * b * c * d * e * e * e * e" - using e2 by (metis comm_mult_left_mono mult_commute linorder_le_cases + using e2 by (metis comm_mult_left_mono mult.commute linorder_le_cases mult_left_mono_neg) qed show "q < r ==> 0 < s ==> s * q < s * r" unfolding less_rat_def - proof (partiality_descending, auto simp add: algebra_simps, simp add: mult_assoc[symmetric]) + proof (partiality_descending, auto simp add: algebra_simps, simp add: mult.assoc[symmetric]) fix a b c d e f :: int assume a: "e \ 0" "f \ 0" "0 \ e * f" "a * b * d * d \ b * b * c * d" have "a * b * d * d * (e * f) \ b * b * c * d * (e * f)" using a by (simp add: mult_right_mono) then show "a * b * d * d * e * f * f * f \ b * b * c * d * e * f * f * f" - by (simp add: mult_assoc[symmetric]) (metis a(3) comm_mult_left_mono - mult_commute mult_left_mono_neg zero_le_mult_iff) + by (simp add: mult.assoc[symmetric]) (metis a(3) comm_mult_left_mono + mult.commute mult_left_mono_neg zero_le_mult_iff) qed show "\z. r \ of_int z" unfolding of_int_rat @@ -258,7 +258,7 @@ fix a b :: int assume "b \ 0" then have "a * b \ (a div b + 1) * b * b" - by (metis mult_commute div_mult_self1_is_id less_int_def linorder_le_cases zdiv_mono1 zdiv_mono1_neg zle_add1_eq_le) + by (metis mult.commute div_mult_self1_is_id less_int_def linorder_le_cases zdiv_mono1 zdiv_mono1_neg zle_add1_eq_le) then show "\z\int. a * b \ z * b * b" by auto qed qed diff -r de51a86fc903 -r cc97b347b301 src/HOL/Rat.thy --- a/src/HOL/Rat.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Rat.thy Fri Jul 04 20:18:47 2014 +0200 @@ -76,7 +76,7 @@ by (simp add: dvd_div_mult_self) with `b \ 0` have "?b \ 0" by auto from `q = Fract a b` `b \ 0` `?b \ 0` have q: "q = Fract ?a ?b" - by (simp add: eq_rat dvd_div_mult mult_commute [of a]) + by (simp add: eq_rat dvd_div_mult mult.commute [of a]) from `b \ 0` have coprime: "coprime ?a ?b" by (auto intro: div_gcd_coprime_int) show C proof (cases "b > 0") @@ -156,7 +156,7 @@ lift_definition inverse_rat :: "rat \ rat" is "\x. if fst x = 0 then (0, 1) else (snd x, fst x)" - by (auto simp add: mult_commute) + by (auto simp add: mult.commute) lemma inverse_rat [simp]: "inverse (Fract a b) = Fract b a" by transfer simp @@ -254,7 +254,7 @@ moreover have "b div gcd a b * gcd a b = b" by (rule dvd_div_mult_self) simp ultimately have "b div gcd a b \ 0" by auto - with False show ?thesis by (simp add: eq_rat dvd_div_mult mult_commute [of a]) + with False show ?thesis by (simp add: eq_rat dvd_div_mult mult.commute [of a]) qed definition normalize :: "int \ int \ int \ int" where @@ -274,7 +274,7 @@ with assms show "p * s = q * r" by (auto simp add: mult_ac sgn_times sgn_0_0) qed from assms show ?thesis - by (auto simp add: normalize_def Let_def dvd_div_div_eq_mult mult_commute sgn_times split: if_splits intro: aux) + by (auto simp add: normalize_def Let_def dvd_div_div_eq_mult mult.commute sgn_times split: if_splits intro: aux) qed lemma normalize_eq: "normalize (a, b) = (p, q) \ Fract p q = Fract a b" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Real.thy --- a/src/HOL/Real.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Real.thy Fri Jul 04 20:18:47 2014 +0200 @@ -412,7 +412,7 @@ lift_definition times_real :: "real \ real \ real" is "\X Y n. X n * Y n" unfolding realrel_def mult_diff_mult - by (subst (4) mult_commute, simp only: cauchy_mult vanishes_add + by (subst (4) mult.commute, simp only: cauchy_mult vanishes_add vanishes_mult_bounded cauchy_imp_bounded simp_thms) lift_definition inverse_real :: "real \ real" @@ -812,7 +812,7 @@ have twos: "\y r :: rat. 0 < r \ \n. y / 2 ^ n < r" apply (simp add: divide_less_eq) - apply (subst mult_commute) + apply (subst mult.commute) apply (frule_tac y=y in ex_less_of_nat_mult) apply clarify apply (rule_tac x=n in exI) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Real_Vector_Spaces.thy Fri Jul 04 20:18:47 2014 +0200 @@ -57,7 +57,7 @@ lemma scale_left_commute: "scale a (scale b x) = scale b (scale a x)" -by (simp add: mult_commute) +by (simp add: mult.commute) lemma scale_zero_left [simp]: "scale 0 x = 0" and scale_minus_left [simp]: "scale (- a) x = - (scale a x)" @@ -255,7 +255,7 @@ by (simp add: of_real_def scaleR_left_diff_distrib) lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y" -by (simp add: of_real_def mult_commute) +by (simp add: of_real_def mult.commute) lemma of_real_setsum[simp]: "of_real (setsum f s) = (\x\s. of_real (f x))" by (induct s rule: infinite_finite_induct) auto @@ -1171,7 +1171,7 @@ lemma sgn_mult: fixes x y :: "'a::real_normed_div_algebra" shows "sgn (x * y) = sgn x * sgn y" -by (simp add: sgn_div_norm norm_mult mult_commute) +by (simp add: sgn_div_norm norm_mult mult.commute) lemma real_sgn_eq: "sgn (x::real) = x / \x\" by (simp add: sgn_div_norm divide_inverse) @@ -1387,7 +1387,7 @@ also have "\ \ (norm x * Kg) * Kf" using g Kf [THEN order_less_imp_le] by (rule mult_right_mono) also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)" - by (rule mult_assoc) + by (rule mult.assoc) finally show "norm (f (g x)) \ norm x * (Kg * Kf)" . qed qed diff -r de51a86fc903 -r cc97b347b301 src/HOL/Rings.thy --- a/src/HOL/Rings.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Rings.thy Fri Jul 04 20:18:47 2014 +0200 @@ -144,7 +144,7 @@ proof - from assms obtain v where "b = a * v" by (auto elim!: dvdE) moreover from assms obtain w where "c = b * w" by (auto elim!: dvdE) - ultimately have "c = a * (v * w)" by (simp add: mult_assoc) + ultimately have "c = a * (v * w)" by (simp add: mult.assoc) then show ?thesis .. qed @@ -160,10 +160,10 @@ by (auto intro!: dvdI) lemma dvd_mult[simp]: "a dvd c \ a dvd (b * c)" -by (auto intro!: mult_left_commute dvdI elim!: dvdE) +by (auto intro!: mult.left_commute dvdI elim!: dvdE) lemma dvd_mult2[simp]: "a dvd b \ a dvd (b * c)" - apply (subst mult_commute) + apply (subst mult.commute) apply (erule dvd_mult) done @@ -185,7 +185,7 @@ qed lemma dvd_mult_left: "a * b dvd c \ a dvd c" -by (simp add: dvd_def mult_assoc, blast) +by (simp add: dvd_def mult.assoc, blast) lemma dvd_mult_right: "a * b dvd c \ b dvd c" unfolding mult_ac [of a] by (rule dvd_mult_left) @@ -684,7 +684,7 @@ fix a b c :: 'a assume "a \ b" "0 \ c" thus "c * a \ c * b" by (rule comm_mult_left_mono) - thus "a * c \ b * c" by (simp only: mult_commute) + thus "a * c \ b * c" by (simp only: mult.commute) qed end @@ -707,7 +707,7 @@ fix a b c :: 'a assume "a < b" "0 < c" thus "c * a < c * b" by (rule comm_mult_strict_left_mono) - thus "a * c < b * c" by (simp only: mult_commute) + thus "a * c < b * c" by (simp only: mult.commute) qed subclass ordered_cancel_comm_semiring diff -r de51a86fc903 -r cc97b347b301 src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Set_Interval.thy Fri Jul 04 20:18:47 2014 +0200 @@ -1480,7 +1480,7 @@ lemma setsum_nat_group: "(\mi\n. f (Suc i)) + f (Suc (Suc n)) = f 0 + ((\i\n. f (Suc i)) + f (Suc (Suc n)))" - by (rule add_assoc) + by (rule add.assoc) also have "(\i\n. f (Suc i)) + f (Suc (Suc n)) = (\i\Suc n. f (Suc i))" by (rule setsum_atMost_Suc [symmetric]) finally show ?case . qed lemma setsum_last_plus: fixes n::nat shows "m <= n \ (\i = m..n. f i) = f n + (\i = m.. 'a::ab_group_add" @@ -1608,7 +1608,7 @@ also from ngt1 have "2*?n*a + d*2*(\i\{1..n - 1}. ?I i) = (2*?n*a + d*?I (n - 1)*?I n)" by (simp only: mult_ac gauss_sum [of "n - 1"], unfold One_nat_def) - (simp add: mult_ac trans [OF add_commute of_nat_Suc [symmetric]]) + (simp add: mult_ac trans [OF add.commute of_nat_Suc [symmetric]]) finally show ?thesis unfolding mult_2 by (simp add: algebra_simps) next diff -r de51a86fc903 -r cc97b347b301 src/HOL/String.thy --- a/src/HOL/String.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/String.thy Fri Jul 04 20:18:47 2014 +0200 @@ -269,7 +269,7 @@ lemma nibble_of_nat_of_char_div_16: "nibble_of_nat (nat_of_char c div 16) = (case c of Char x y \ x)" by (cases c) - (simp add: nat_of_char_def add_commute nat_of_nibble_less_16) + (simp add: nat_of_char_def add.commute nat_of_nibble_less_16) lemma nibble_of_nat_nat_of_char: "nibble_of_nat (nat_of_char c) = (case c of Char x y \ y)" @@ -279,7 +279,7 @@ by (simp add: nibble_of_nat_mod_16) then have "nibble_of_nat (nat_of_nibble y + nat_of_nibble x * 16) = y" by (simp only: nibble_of_nat_mod_16) - with Char show ?thesis by (simp add: nat_of_char_def add_commute) + with Char show ?thesis by (simp add: nat_of_char_def add.commute) qed code_datatype Char -- {* drop case certificate for char *} @@ -303,7 +303,7 @@ have "n mod 256 = 16 * (n div 16 mod 16) + n mod 16" by simp then show ?thesis by (simp add: char_of_nat_def enum_char_product_enum_nibble card_UNIV_nibble - card_UNIV_length_enum [symmetric] nibble_of_nat_def product_nth add_commute) + card_UNIV_length_enum [symmetric] nibble_of_nat_def product_nth add.commute) qed lemma char_of_nat_of_char [simp]: diff -r de51a86fc903 -r cc97b347b301 src/HOL/Tools/numeral_simprocs.ML --- a/src/HOL/Tools/numeral_simprocs.ML Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Tools/numeral_simprocs.ML Fri Jul 04 20:18:47 2014 +0200 @@ -227,7 +227,7 @@ (*combine unary minus with numeric literals, however nested within a product*) val mult_minus_simps = - [@{thm mult_assoc}, @{thm minus_mult_right}, @{thm minus_mult_commute}]; + [@{thm mult.assoc}, @{thm minus_mult_right}, @{thm minus_mult_commute}]; val norm_ss1 = simpset_of (put_simpset num_ss @{context} @@ -719,7 +719,7 @@ @{thm diff_conv_add_uminus}, @{thm "minus_divide_left"}, @{thm "add_divide_distrib"} RS sym, @{thm field_divide_inverse} RS sym, @{thm inverse_divide}, - Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.rewr_conv (mk_meta_eq @{thm mult_commute})))) + Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.rewr_conv (mk_meta_eq @{thm mult.commute})))) (@{thm field_divide_inverse} RS sym)] val field_comp_ss = diff -r de51a86fc903 -r cc97b347b301 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Transcendental.thy Fri Jul 04 20:18:47 2014 +0200 @@ -61,7 +61,7 @@ also have "... = y * ((x - y) * (\ppp = K * norm (z ^ n) * (inverse (norm (x ^ n)) * norm (x ^ n))" by (simp add: x_neq_0) also have "\ = K * norm (z ^ n) * inverse (norm (x ^ n)) * norm (x ^ n)" - by (simp only: mult_assoc) + by (simp only: mult.assoc) finally show "norm (norm (f n * z ^ n)) \ K * norm (z ^ n) * inverse (norm (x ^ n))" by (simp add: mult_le_cancel_right x_neq_0) @@ -145,7 +145,7 @@ thus "summable (\n. K * norm (z ^ n) * inverse (norm (x ^ n)))" using x_neq_0 by (simp add: norm_mult nonzero_norm_inverse power_mult_distrib - power_inverse norm_power mult_assoc) + power_inverse norm_power mult.assoc) qed ultimately show "summable (\n. norm (f n * z ^ n))" by (rule summable_comparison_test) @@ -469,10 +469,10 @@ (z + h) ^ q * z ^ (n - 2 - q))" (is "?lhs = ?rhs") apply (subgoal_tac "h * ?lhs = h * ?rhs", simp add: h) apply (simp add: right_diff_distrib diff_divide_distrib h) - apply (simp add: mult_assoc [symmetric]) + apply (simp add: mult.assoc [symmetric]) apply (cases "n", simp) apply (simp add: lemma_realpow_diff_sumr2 h - right_diff_distrib [symmetric] mult_assoc + right_diff_distrib [symmetric] mult.assoc del: power_Suc setsum_lessThan_Suc of_nat_Suc) apply (subst lemma_realpow_rev_sumr) apply (subst sumr_diff_mult_const2) @@ -483,7 +483,7 @@ apply (clarify) apply (simp add: setsum_right_distrib lemma_realpow_diff_sumr2 mult_ac del: setsum_lessThan_Suc power_Suc) - apply (subst mult_assoc [symmetric], subst power_add [symmetric]) + apply (subst mult.assoc [symmetric], subst power_add [symmetric]) apply (simp add: mult_ac) done @@ -508,7 +508,7 @@ have "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) = norm (\pq \ of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2)) * norm h" proof (rule mult_right_mono [OF _ norm_ge_zero]) from norm_ge_zero 2 have K: "0 \ K" @@ -530,7 +530,7 @@ done qed also have "\ = of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" - by (simp only: mult_assoc) + by (simp only: mult.assoc) finally show ?thesis . qed @@ -638,9 +638,9 @@ by (rule order_le_less_trans) show "norm (c n * (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0))) \ norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2) * norm h" - apply (simp only: norm_mult mult_assoc) + apply (simp only: norm_mult mult.assoc) apply (rule mult_left_mono [OF _ norm_ge_zero]) - apply (simp add: mult_assoc [symmetric]) + apply (simp add: mult.assoc [symmetric]) apply (metis h lemma_termdiff3 less_eq_real_def r1 xh) done qed @@ -797,7 +797,7 @@ also have "\ \ ?diff_part + \ (\n. ?diff (n + ?N) x) - (\ n. f' x0 (n + ?N)) \" unfolding suminf_split_initial_segment[OF all_smbl, where k="?N"] unfolding suminf_diff[OF div_shft_smbl ign[OF `summable (f' x0)`]] - apply (subst (5) add_commute) + apply (subst (5) add.commute) by (rule abs_triangle_ineq) also have "\ \ ?diff_part + ?L_part + ?f'_part" using abs_triangle_ineq4 by auto @@ -877,7 +877,7 @@ unfolding abs_mult[symmetric] by auto qed also have "\ = \f n * real (Suc n) * R' ^ n\ * \x - y\" - unfolding abs_mult mult_assoc[symmetric] by algebra + unfolding abs_mult mult.assoc[symmetric] by algebra finally show ?thesis . qed } @@ -900,7 +900,7 @@ unfolding real_norm_def abs_mult by (rule mult_right_mono) (auto simp add: le[unfolded mult_1_right]) qed (rule powser_insidea[OF converges[OF `R' \ {-R <..< R}`] `norm x < norm R'`]) - from this[THEN summable_mult2[where c=x], unfolded mult_assoc, unfolded mult_commute] + from this[THEN summable_mult2[where c=x], unfolded mult.assoc, unfolded mult.commute] show "summable (?f x)" by auto } show "summable (?f' x0)" @@ -984,7 +984,7 @@ lemma exp_fdiffs: "diffs (\n. inverse(real (fact n))) = (\n. inverse(real (fact n)))" - by (simp add: diffs_def mult_assoc [symmetric] real_of_nat_def of_nat_mult + by (simp add: diffs_def mult.assoc [symmetric] real_of_nat_def of_nat_mult del: mult_Suc of_nat_Suc) lemma diffs_of_real: "diffs (\n. of_real (f n)) = (\n. of_real (diffs f n))" @@ -1132,7 +1132,7 @@ by simp lemma exp_real_of_nat_mult: "exp(real n * x) = exp(x) ^ n" - by (induct n) (auto simp add: real_of_nat_Suc distrib_left exp_add mult_commute) + by (induct n) (auto simp add: real_of_nat_Suc distrib_left exp_add mult.commute) text {* Strict monotonicity of exponential. *} @@ -1824,7 +1824,7 @@ by (simp add: powr_def) lemma powr_powr_swap: "(x powr a) powr b = (x powr b) powr a" - by (simp add: powr_powr mult_commute) + by (simp add: powr_powr mult.commute) lemma powr_minus: "x powr (-a) = inverse (x powr a)" by (simp add: powr_def exp_minus [symmetric]) @@ -2006,7 +2006,7 @@ by(simp add: root_powr_inverse ln_powr) lemma ln_sqrt: "0 < x \ ln (sqrt x) = ln x / 2" - by (simp add: ln_powr powr_numeral ln_powr[symmetric] mult_commute) + by (simp add: ln_powr powr_numeral ln_powr[symmetric] mult.commute) lemma log_root: "\ n > 0; a > 0 \ \ log b (root n a) = log b a / n" by(simp add: log_def ln_root) @@ -2078,7 +2078,7 @@ unfolding powr_def exp_inj_iff by simp lemma ln_powr_bound: "1 <= x ==> 0 < a ==> ln x <= (x powr a) / a" - by (metis less_eq_real_def ln_less_self mult_imp_le_div_pos ln_powr mult_commute + by (metis less_eq_real_def ln_less_self mult_imp_le_div_pos ln_powr mult.commute order.strict_trans2 powr_gt_zero zero_less_one) lemma ln_powr_bound2: @@ -2339,7 +2339,7 @@ qed lemma sin_cos_squared_add2 [simp]: "(cos x)\<^sup>2 + (sin x)\<^sup>2 = 1" - by (subst add_commute, rule sin_cos_squared_add) + by (subst add.commute, rule sin_cos_squared_add) lemma sin_cos_squared_add3 [simp]: "cos x * cos x + sin x * sin x = 1" using sin_cos_squared_add2 [unfolded power2_eq_square] . @@ -2422,13 +2422,13 @@ using sin_add [of x "- y"] by simp lemma sin_diff2: "sin (x - y) = cos y * sin x - sin y * cos x" - by (simp add: sin_diff mult_commute) + by (simp add: sin_diff mult.commute) lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y" using cos_add [of x "- y"] by simp lemma cos_diff2: "cos (x - y) = cos y * cos x + sin y * sin x" - by (simp add: cos_diff mult_commute) + by (simp add: cos_diff mult.commute) lemma sin_double [simp]: "sin(2 * x) = 2* sin x * cos x" using sin_add [where x=x and y=x] by simp @@ -2567,7 +2567,7 @@ note *** = this have [simp]: "\x y::real. 0 < x - y \ y < x" by arith from ** show ?thesis by (rule sumr_pos_lt_pair) - (simp add: divide_inverse mult_assoc [symmetric] ***) + (simp add: divide_inverse mult.assoc [symmetric] ***) qed ultimately have "0 < (\n. - (-1 ^ n * 2 ^ (2 * n) / real (fact (2 * n))))" by (rule order_less_trans) @@ -2681,13 +2681,13 @@ by (induct n) (auto simp add: real_of_nat_Suc distrib_right) lemma cos_npi2 [simp]: "cos (pi * real n) = -1 ^ n" - by (metis cos_npi mult_commute) + by (metis cos_npi mult.commute) lemma sin_npi [simp]: "sin (real (n::nat) * pi) = 0" by (induct n) (auto simp add: real_of_nat_Suc distrib_right) lemma sin_npi2 [simp]: "sin (pi * real (n::nat)) = 0" - by (simp add: mult_commute [of pi]) + by (simp add: mult.commute [of pi]) lemma cos_two_pi [simp]: "cos (2 * pi) = 1" by (simp add: cos_double) @@ -2785,7 +2785,7 @@ apply (cut_tac y="-y" in cos_total, simp) apply simp apply (erule ex1E) apply (rule_tac a = "x - (pi/2)" in ex1I) -apply (simp (no_asm) add: add_assoc) +apply (simp (no_asm) add: add.assoc) apply (rotate_tac 3) apply (drule_tac x = "xa + pi/2" in spec, safe, simp_all add: cos_add) done @@ -3023,7 +3023,7 @@ apply (auto simp add: divide_inverse) apply (rule mult_pos_pos) apply (subgoal_tac [3] "0 < sin e & 0 < cos e") - apply (auto intro: cos_gt_zero sin_gt_zero2 simp add: mult_commute) + apply (auto intro: cos_gt_zero sin_gt_zero2 simp add: mult.commute) done lemma tan_total_pos: "0 \ y ==> \x. 0 \ x & x < pi/2 & tan x = y" @@ -3321,7 +3321,7 @@ apply (rule_tac c1 = "(cos x)\<^sup>2" in mult_right_cancel [THEN iffD1]) apply (auto dest: field_power_not_zero simp add: power_mult_distrib distrib_right power_divide tan_def - mult_assoc power_inverse [symmetric]) + mult.assoc power_inverse [symmetric]) done lemma arctan_less_iff: "arctan x < arctan y \ x < y" @@ -3559,11 +3559,11 @@ by (auto simp add: algebra_simps sin_add) thus ?thesis by (simp add: real_of_nat_Suc distrib_right add_divide_distrib - mult_commute [of pi]) + mult.commute [of pi]) qed lemma cos_2npi [simp]: "cos (2 * real (n::nat) * pi) = 1" - by (simp add: cos_double mult_assoc power_add [symmetric] numeral_2_eq_2) + by (simp add: cos_double mult.assoc power_add [symmetric] numeral_2_eq_2) lemma cos_3over2_pi [simp]: "cos (3 / 2 * pi) = 0" apply (subgoal_tac "cos (pi + pi/2) = 0", simp) @@ -3571,7 +3571,7 @@ done lemma sin_2npi [simp]: "sin (2 * real (n::nat) * pi) = 0" - by (auto simp add: mult_assoc) + by (auto simp add: mult.assoc) lemma sin_3over2_pi [simp]: "sin (3 / 2 * pi) = - 1" apply (subgoal_tac "sin (pi + pi/2) = - 1", simp) @@ -3771,14 +3771,14 @@ next fix x :: real assume "(\ n. if even n then f (n div 2) else 0) sums x" - from LIMSEQ_linear[OF this[unfolded sums_def] pos2, unfolded sum_split_even_odd[unfolded mult_commute]] + from LIMSEQ_linear[OF this[unfolded sums_def] pos2, unfolded sum_split_even_odd[unfolded mult.commute]] show "f sums x" unfolding sums_def by auto qed hence "op sums f = op sums (\ n. if even n then f (n div 2) else 0)" .. } note sums_even = this have Int_eq: "(\n. ?f n * real (Suc n) * x^n) = ?Int" - unfolding if_eq mult_commute[of _ 2] suminf_def sums_even[of "\ n. -1 ^ n * x ^ (2 * n)", symmetric] + unfolding if_eq mult.commute[of _ 2] suminf_def sums_even[of "\ n. -1 ^ n * x ^ (2 * n)", symmetric] by auto { @@ -3858,7 +3858,7 @@ hence "(\ n. (- (x\<^sup>2)) ^ n) sums (1 / (1 - (- (x\<^sup>2))))" unfolding real_norm_def[symmetric] by (rule geometric_sums) hence "(?c' x) sums (1 / (1 - (- (x\<^sup>2))))" - unfolding power_mult_distrib[symmetric] power_mult nat_mult_commute[of _ 2] by auto + unfolding power_mult_distrib[symmetric] power_mult mult.commute[of _ 2] by auto hence suminf_c'_eq_geom: "inverse (1 + x\<^sup>2) = suminf (?c' x)" using sums_unique unfolding inverse_eq_divide by auto have "DERIV (\ x. suminf (?c x)) x :> (inverse (1 + x\<^sup>2))" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Word/Misc_Numeric.thy --- a/src/HOL/Word/Misc_Numeric.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Word/Misc_Numeric.thy Fri Jul 04 20:18:47 2014 +0200 @@ -25,7 +25,7 @@ lemma emep1: fixes n d :: int shows "even n \ even d \ 0 \ d \ (n + 1) mod d = (n mod d) + 1" - by (auto simp add: pos_zmod_mult_2 add_commute even_equiv_def) + by (auto simp add: pos_zmod_mult_2 add.commute even_equiv_def) lemma int_mod_ge: "a < n \ 0 < (n :: int) \ a \ a mod n" @@ -63,7 +63,7 @@ lemma p1mod22k: fixes b :: int shows "(2 * b + 1) mod (2 * 2 ^ n) = 2 * (b mod 2 ^ n) + 1" - by (simp add: p1mod22k' add_commute) + by (simp add: p1mod22k' add.commute) lemma int_mod_lem: "(0 :: int) < n ==> (0 <= b & b < n) = (b mod n = b)" diff -r de51a86fc903 -r cc97b347b301 src/HOL/Word/Word.thy --- a/src/HOL/Word/Word.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Word/Word.thy Fri Jul 04 20:18:47 2014 +0200 @@ -1753,7 +1753,7 @@ prefer 2 apply (insert uint_range' [of s])[1] apply arith - apply (drule add_commute [THEN xtr1]) + apply (drule add.commute [THEN xtr1]) apply (simp add: diff_less_eq [symmetric]) apply (drule less_le_mult) apply arith @@ -2161,7 +2161,7 @@ "a div b * b \ (a::nat)" using gt_or_eq_0 [of b] apply (rule disjE) - apply (erule xtr4 [OF thd mult_commute]) + apply (erule xtr4 [OF thd mult.commute]) apply clarsimp done @@ -2921,7 +2921,7 @@ apply (unfold shiftr_def) apply (induct "n") apply simp - apply (simp add: shiftr1_div_2 mult_commute + apply (simp add: shiftr1_div_2 mult.commute zdiv_zmult2_eq [symmetric]) done @@ -2929,7 +2929,7 @@ apply (unfold sshiftr_def) apply (induct "n") apply simp - apply (simp add: sshiftr1_div_2 mult_commute + apply (simp add: sshiftr1_div_2 mult.commute zdiv_zmult2_eq [symmetric]) done @@ -3731,7 +3731,7 @@ apply (clarsimp simp: word_ubin.eq_norm nth_bintr word_size split: prod.splits) apply (drule bin_nth_split) apply safe - apply (simp_all add: add_commute) + apply (simp_all add: add.commute) apply (erule bin_nth_uint_imp)+ done @@ -3895,7 +3895,7 @@ lemma foldl_eq_foldr: "foldl op + x xs = foldr op + (x # xs) (0 :: 'a :: comm_monoid_add)" - by (induct xs arbitrary: x) (auto simp add : add_assoc) + by (induct xs arbitrary: x) (auto simp add : add.assoc) lemmas test_bit_cong = arg_cong [where f = "test_bit", THEN fun_cong] @@ -3935,7 +3935,7 @@ (* alternative proof of word_rcat_rsplit *) lemmas tdle = iffD2 [OF split_div_lemma refl, THEN conjunct1] -lemmas dtle = xtr4 [OF tdle mult_commute] +lemmas dtle = xtr4 [OF tdle mult.commute] lemma word_rcat_rsplit: "word_rcat (word_rsplit w) = w" apply (rule word_eqI) @@ -3959,7 +3959,7 @@ fixes n::nat shows "0 < n \ (k * n + m) div n = m div n + k" and "(k * n + m) mod n = m mod n" - by (auto simp: add_commute) + by (auto simp: add.commute) lemma word_rsplit_rcat_size [OF refl]: "word_rcat (ws :: 'a :: len word list) = frcw \ diff -r de51a86fc903 -r cc97b347b301 src/HOL/Word/WordBitwise.thy --- a/src/HOL/Word/WordBitwise.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Word/WordBitwise.thy Fri Jul 04 20:18:47 2014 +0200 @@ -306,7 +306,7 @@ show ?case using Cons - apply (simp add: trans [OF of_bl_append add_commute] + apply (simp add: trans [OF of_bl_append add.commute] rbl_mul_simps rbl_word_plus' Cons.hyps distrib_right mult_bit shiftl rbl_shiftl) diff -r de51a86fc903 -r cc97b347b301 src/HOL/Word/Word_Miscellaneous.thy --- a/src/HOL/Word/Word_Miscellaneous.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/Word/Word_Miscellaneous.thy Fri Jul 04 20:18:47 2014 +0200 @@ -146,14 +146,14 @@ lemma emep1: "even n ==> even d ==> 0 <= d ==> (n + 1) mod (d :: int) = (n mod d) + 1" - apply (simp add: add_commute) + apply (simp add: add.commute) apply (safe dest!: even_equiv_def [THEN iffD1]) apply (subst pos_zmod_mult_2) apply arith apply (simp add: mod_mult_mult1) done -lemmas eme1p = emep1 [simplified add_commute] +lemmas eme1p = emep1 [simplified add.commute] lemma le_diff_eq': "(a \ c - b) = (b + a \ (c::int))" by arith @@ -187,9 +187,9 @@ lemmas iszero_minus = trans [THEN trans, OF iszero_def neg_equal_0_iff_equal iszero_def [symmetric]] -lemmas zadd_diff_inverse = trans [OF diff_add_cancel [symmetric] add_commute] +lemmas zadd_diff_inverse = trans [OF diff_add_cancel [symmetric] add.commute] -lemmas add_diff_cancel2 = add_commute [THEN diff_eq_eq [THEN iffD2]] +lemmas add_diff_cancel2 = add.commute [THEN diff_eq_eq [THEN iffD2]] lemmas rdmods [symmetric] = mod_minus_eq mod_diff_left_eq mod_diff_right_eq mod_add_left_eq @@ -296,13 +296,13 @@ "a + b = c + d ==> a >= c & b <= d | a <= c & b >= (d :: nat)" by arith -lemmas pl_pl_rels' = add_commute [THEN [2] trans, THEN pl_pl_rels] +lemmas pl_pl_rels' = add.commute [THEN [2] trans, THEN pl_pl_rels] lemma minus_eq: "(m - k = m) = (k = 0 | m = (0 :: nat))" by arith lemma pl_pl_mm: "(a :: nat) + b = c + d ==> a - c = d - b" by arith -lemmas pl_pl_mm' = add_commute [THEN [2] trans, THEN pl_pl_mm] +lemmas pl_pl_mm' = add.commute [THEN [2] trans, THEN pl_pl_mm] lemmas dme = box_equals [OF div_mod_equality add_0_right add_0_right] lemmas dtle = xtr3 [OF dme [symmetric] le_add1] @@ -341,7 +341,7 @@ apply clarsimp apply clarify apply (cases "b > 0") - apply (drule mult_commute [THEN xtr1]) + apply (drule mult.commute [THEN xtr1]) apply (frule (1) td_gal_lt [THEN iffD1]) apply (clarsimp simp: le_simps) apply (rule mult_div_cancel [THEN [2] xtr4]) diff -r de51a86fc903 -r cc97b347b301 src/HOL/ex/Adhoc_Overloading_Examples.thy --- a/src/HOL/ex/Adhoc_Overloading_Examples.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/ex/Adhoc_Overloading_Examples.thy Fri Jul 04 20:18:47 2014 +0200 @@ -211,7 +211,7 @@ apply default unfolding permute_perm_def apply simp - apply (simp only: diff_conv_add_uminus minus_add add_assoc) + apply (simp only: diff_conv_add_uminus minus_add add.assoc) done text {*Permuting functions.*} diff -r de51a86fc903 -r cc97b347b301 src/HOL/ex/Dedekind_Real.thy --- a/src/HOL/ex/Dedekind_Real.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/ex/Dedekind_Real.thy Fri Jul 04 20:18:47 2014 +0200 @@ -229,7 +229,7 @@ lemma preal_add_commute: "(x::preal) + y = y + x" apply (unfold preal_add_def add_set_def) apply (rule_tac f = Abs_preal in arg_cong) -apply (force simp add: add_commute) +apply (force simp add: add.commute) done text{*Lemmas for proving that addition of two positive reals gives @@ -344,7 +344,7 @@ lemma preal_mult_commute: "(x::preal) * y = y * x" apply (unfold preal_mult_def mult_set_def) apply (rule_tac f = Abs_preal in arg_cong) -apply (force simp add: mult_commute) +apply (force simp add: mult.commute) done text{*Multiplication of two positive reals gives a positive real.*} @@ -420,7 +420,7 @@ show "\y'\B. z = (z/y) * y'" proof show "z = (z/y)*y" - by (simp add: divide_inverse mult_commute [of y] mult_assoc + by (simp add: divide_inverse mult.commute [of y] mult.assoc order_less_imp_not_eq2) show "y \ B" by fact qed @@ -501,7 +501,7 @@ show "\v'\A. x = (x / v) * v'" proof show "x = (x/v)*v" - by (simp add: divide_inverse mult_assoc vpos + by (simp add: divide_inverse mult.assoc vpos order_less_imp_not_eq2) show "v \ A" by fact qed @@ -550,7 +550,7 @@ have cpos: "0 < ?c" by (simp add: zero_less_divide_iff zero_less_mult_iff pos_add_strict) show "a * d + b * e = ?c * (d + e)" - by (simp add: divide_inverse mult_assoc order_less_imp_not_eq2) + by (simp add: divide_inverse mult.assoc order_less_imp_not_eq2) show "?c \ Rep_preal w" proof (cases rule: linorder_le_cases) assume "a \ b" @@ -777,7 +777,7 @@ qed hence "y < r" by simp with ypos have dless: "?d < (r * ?d)/y" - by (simp add: pos_less_divide_eq mult_commute [of ?d] + by (simp add: pos_less_divide_eq mult.commute [of ?d] mult_strict_right_mono dpos) have "r + ?d < r*x" proof - @@ -826,10 +826,10 @@ show "x/u < x/r" using xpos upos rpos by (simp add: divide_inverse mult_less_cancel_left rless) show "inverse (x / r) \ Rep_preal R" using notin - by (simp add: divide_inverse mult_commute) + by (simp add: divide_inverse mult.commute) show "u \ Rep_preal R" by (rule u) show "x = x / u * u" using upos - by (simp add: divide_inverse mult_commute) + by (simp add: divide_inverse mult.commute) qed qed @@ -1284,9 +1284,9 @@ lemma real_add_congruent2_lemma: "[|a + ba = aa + b; ab + bc = ac + bb|] ==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))" -apply (simp add: add_assoc) -apply (rule add_left_commute [of ab, THEN ssubst]) -apply (simp add: add_assoc [symmetric]) +apply (simp add: add.assoc) +apply (rule add.left_commute [of ab, THEN ssubst]) +apply (simp add: add.assoc [symmetric]) apply (simp add: add_ac) done @@ -1305,7 +1305,7 @@ lemma real_minus: "- Abs_Real(realrel``{(x,y)}) = Abs_Real(realrel `` {(y,x)})" proof - have "(\(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel" - by (auto simp add: congruent_def add_commute) + by (auto simp add: congruent_def add.commute) thus ?thesis by (simp add: real_minus_def UN_equiv_class [OF equiv_realrel]) qed @@ -1314,13 +1314,13 @@ proof fix x y z :: real show "(x + y) + z = x + (y + z)" - by (cases x, cases y, cases z, simp add: real_add add_assoc) + by (cases x, cases y, cases z, simp add: real_add add.assoc) show "x + y = y + x" - by (cases x, cases y, simp add: real_add add_commute) + by (cases x, cases y, simp add: real_add add.commute) show "0 + x = x" by (cases x, simp add: real_add real_zero_def add_ac) show "- x + x = 0" - by (cases x, simp add: real_minus real_add real_zero_def add_commute) + by (cases x, simp add: real_minus real_add real_zero_def add.commute) show "x - y = x + - y" by (simp add: real_diff_def) qed @@ -1332,9 +1332,9 @@ "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==> x * x1 + y * y1 + (x * y2 + y * x2) = x * x2 + y * y2 + (x * y1 + y * x1)" -apply (simp add: add_left_commute add_assoc [symmetric]) -apply (simp add: add_assoc distrib_left [symmetric]) -apply (simp add: add_commute) +apply (simp add: add.left_commute add.assoc [symmetric]) +apply (simp add: add.assoc distrib_left [symmetric]) +apply (simp add: add.commute) done lemma real_mult_congruent2: @@ -1343,7 +1343,7 @@ { Abs_Real (realrel``{(x1*x2 + y1*y2, x1*y2+y1*x2)}) }) p2) p1) respects2 realrel" apply (rule congruent2_commuteI [OF equiv_realrel], clarify) -apply (simp add: mult_commute add_commute) +apply (simp add: mult.commute add.commute) apply (auto simp add: real_mult_congruent2_lemma) done @@ -1393,7 +1393,7 @@ subsection {* Inverse and Division *} lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0" -by (simp add: real_zero_def add_commute) +by (simp add: real_zero_def add.commute) text{*Instead of using an existential quantifier and constructing the inverse within the proof, we could define the inverse explicitly.*} @@ -1416,8 +1416,8 @@ apply (drule real_mult_inverse_left_ex, safe) apply (rule theI, assumption, rename_tac z) apply (subgoal_tac "(z * x) * y = z * (x * y)") -apply (simp add: mult_commute) -apply (rule mult_assoc) +apply (simp add: mult.commute) +apply (rule mult.assoc) done @@ -1510,7 +1510,7 @@ apply (cases x, cases y) apply (auto simp add: real_le real_zero_def real_diff_def real_add real_minus add_ac) -apply (simp_all add: add_assoc [symmetric]) +apply (simp_all add: add.assoc [symmetric]) done lemma real_add_left_mono: @@ -1591,7 +1591,7 @@ apply (simp add: real_of_preal_def real_zero_def, cases x) apply (auto simp add: real_minus add_ac) apply (cut_tac x = xa and y = y in linorder_less_linear) -apply (auto dest!: less_add_left_Ex simp add: add_assoc [symmetric]) +apply (auto dest!: less_add_left_Ex simp add: add.assoc [symmetric]) done lemma real_of_preal_leD: diff -r de51a86fc903 -r cc97b347b301 src/HOL/ex/Gauge_Integration.thy --- a/src/HOL/ex/Gauge_Integration.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/ex/Gauge_Integration.thy Fri Jul 04 20:18:47 2014 +0200 @@ -510,9 +510,9 @@ next case False then have "inverse (z - x) * (f z - f x - f' x * (z - x)) = (f z - f x) / (z - x) - f' x" - apply (subst mult_commute) + apply (subst mult.commute) apply (simp add: left_diff_distrib) - apply (simp add: mult_assoc divide_inverse) + apply (simp add: mult.assoc divide_inverse) apply (simp add: ring_distribs) done moreover from False `\z - x\ < s` have "\(f z - f x) / (z - x) - f' x\ < e / 2" diff -r de51a86fc903 -r cc97b347b301 src/HOL/ex/ThreeDivides.thy --- a/src/HOL/ex/ThreeDivides.thy Fri Jul 04 20:07:08 2014 +0200 +++ b/src/HOL/ex/ThreeDivides.thy Fri Jul 04 20:18:47 2014 +0200 @@ -88,7 +88,7 @@ lemma three_divs_1: fixes D :: "nat \ nat" shows "3 dvd (\x