--- a/NEWS Thu Oct 31 16:54:22 2013 +0100
+++ b/NEWS Fri Nov 01 18:51:14 2013 +0100
@@ -15,6 +15,39 @@
even_zero_(nat|int) ~> even_zero
INCOMPATIBILITY.
+*** HOL ***
+
+* Elimination of fact duplicates:
+ equals_zero_I ~> minus_unique
+ diff_eq_0_iff_eq ~> right_minus_eq
+INCOMPATIBILITY.
+
+* Fact name consolidation:
+ diff_def, diff_minus, ab_diff_minus ~> diff_conv_add_uminus
+INCOMPATIBILITY.
+
+* More simplification rules on unary and binary minus:
+add_diff_cancel, add_diff_cancel_left, add_le_same_cancel1,
+add_le_same_cancel2, add_less_same_cancel1, add_less_same_cancel2,
+add_minus_cancel, diff_add_cancel, le_add_same_cancel1,
+le_add_same_cancel2, less_add_same_cancel1, less_add_same_cancel2,
+minus_add_cancel, uminus_add_conv_diff. These correspondingly
+have been taken away from fact collections algebra_simps and
+field_simps. INCOMPATIBILITY.
+
+To restore proofs, the following patterns are helpful:
+
+a) Arbitrary failing proof not involving "diff_def":
+Consider simplification with algebra_simps or field_simps.
+
+b) Lifting rules from addition to subtraction:
+Try with "using <rule for addition> of [… "- _" …]" by simp".
+
+c) Simplification with "diff_def": just drop "diff_def".
+Consider simplification with algebra_simps or field_simps;
+or the brute way with
+"simp add: diff_conv_add_uminus del: add_uminus_conv_diff".
+
New in Isabelle2013-1 (November 2013)
-------------------------------------
--- a/src/HOL/Big_Operators.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Big_Operators.thy Fri Nov 01 18:51:14 2013 +0100
@@ -696,11 +696,7 @@
lemma setsum_subtractf:
"setsum (%x. ((f x)::'a::ab_group_add) - g x) A =
setsum f A - setsum g A"
-proof (cases "finite A")
- case True thus ?thesis by (simp add: diff_minus setsum_addf setsum_negf)
-next
- case False thus ?thesis by simp
-qed
+ using setsum_addf [of f "- g" A] by (simp add: setsum_negf)
lemma setsum_nonneg:
assumes nn: "\<forall>x\<in>A. (0::'a::{ordered_ab_semigroup_add,comm_monoid_add}) \<le> f x"
--- a/src/HOL/Complex.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Complex.thy Fri Nov 01 18:51:14 2013 +0100
@@ -587,7 +587,7 @@
by (simp add: cis_def)
lemma cis_divide: "cis a / cis b = cis (a - b)"
- by (simp add: complex_divide_def cis_mult diff_minus)
+ by (simp add: complex_divide_def cis_mult)
lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re(cis a ^ n)"
by (auto simp add: DeMoivre)
--- a/src/HOL/Decision_Procs/Approximation.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy Fri Nov 01 18:51:14 2013 +0100
@@ -29,7 +29,7 @@
have shift_pow: "\<And>i. - (x * ((-1)^i * a (Suc i) * x ^ i)) = (-1)^(Suc i) * a (Suc i) * x ^ (Suc i)"
by auto
show ?thesis
- unfolding setsum_right_distrib shift_pow diff_minus setsum_negf[symmetric]
+ unfolding setsum_right_distrib shift_pow uminus_add_conv_diff [symmetric] setsum_negf[symmetric]
setsum_head_upt_Suc[OF zero_less_Suc]
setsum_reindex[OF inj_Suc, unfolded comp_def, symmetric, of "\<lambda> n. (-1)^n *a n * x^n"] by auto
qed
@@ -533,12 +533,12 @@
have "pi \<le> ub_pi n"
unfolding ub_pi_def machin_pi Let_def unfolding Float_num
using lb_arctan[of 239] ub_arctan[of 5] powr_realpow[of 2 2]
- by (auto intro!: mult_left_mono add_mono simp add: diff_minus)
+ by (auto intro!: mult_left_mono add_mono simp add: uminus_add_conv_diff [symmetric] simp del: uminus_add_conv_diff)
moreover
have "lb_pi n \<le> pi"
unfolding lb_pi_def machin_pi Let_def Float_num
using lb_arctan[of 5] ub_arctan[of 239] powr_realpow[of 2 2]
- by (auto intro!: mult_left_mono add_mono simp add: diff_minus)
+ by (auto intro!: mult_left_mono add_mono simp add: uminus_add_conv_diff [symmetric] simp del: uminus_add_conv_diff)
ultimately show ?thesis by auto
qed
@@ -1208,8 +1208,8 @@
using x unfolding k[symmetric]
by (cases "k = 0")
(auto intro!: add_mono
- simp add: diff_minus k[symmetric]
- simp del: float_of_numeral)
+ simp add: k [symmetric] uminus_add_conv_diff [symmetric]
+ simp del: float_of_numeral uminus_add_conv_diff)
note lx = this[THEN conjunct1] and ux = this[THEN conjunct2]
hence lx_less_ux: "?lx \<le> real ?ux" by (rule order_trans)
@@ -1223,7 +1223,7 @@
also have "\<dots> \<le> cos (x + (-k) * (2 * pi))"
using cos_monotone_minus_pi_0'[OF pi_lx lx x_le_0]
by (simp only: uminus_float.rep_eq real_of_int_minus
- cos_minus diff_minus mult_minus_left)
+ cos_minus mult_minus_left) simp
finally have "(lb_cos prec (- ?lx)) \<le> cos x"
unfolding cos_periodic_int . }
note negative_lx = this
@@ -1236,7 +1236,7 @@
have "cos (x + (-k) * (2 * pi)) \<le> cos ?lx"
using cos_monotone_0_pi'[OF lx_0 lx pi_x]
by (simp only: real_of_int_minus
- cos_minus diff_minus mult_minus_left)
+ cos_minus mult_minus_left) simp
also have "\<dots> \<le> (ub_cos prec ?lx)"
using lb_cos[OF lx_0 pi_lx] by simp
finally have "cos x \<le> (ub_cos prec ?lx)"
@@ -1251,7 +1251,7 @@
have "cos (x + (-k) * (2 * pi)) \<le> cos (real (- ?ux))"
using cos_monotone_minus_pi_0'[OF pi_x ux ux_0]
by (simp only: uminus_float.rep_eq real_of_int_minus
- cos_minus diff_minus mult_minus_left)
+ cos_minus mult_minus_left) simp
also have "\<dots> \<le> (ub_cos prec (- ?ux))"
using lb_cos_minus[OF pi_ux ux_0, of prec] by simp
finally have "cos x \<le> (ub_cos prec (- ?ux))"
@@ -1268,7 +1268,7 @@
also have "\<dots> \<le> cos (x + (-k) * (2 * pi))"
using cos_monotone_0_pi'[OF x_ge_0 ux pi_ux]
by (simp only: real_of_int_minus
- cos_minus diff_minus mult_minus_left)
+ cos_minus mult_minus_left) simp
finally have "(lb_cos prec ?ux) \<le> cos x"
unfolding cos_periodic_int . }
note positive_ux = this
@@ -1343,7 +1343,7 @@
also have "\<dots> \<le> cos ((?ux - 2 * ?lpi))"
using cos_monotone_minus_pi_0'[OF pi_x x_le_ux ux_0]
by (simp only: minus_float.rep_eq real_of_int_minus real_of_one minus_one[symmetric]
- diff_minus mult_minus_left mult_1_left)
+ mult_minus_left mult_1_left) simp
also have "\<dots> = cos ((- (?ux - 2 * ?lpi)))"
unfolding uminus_float.rep_eq cos_minus ..
also have "\<dots> \<le> (ub_cos prec (- (?ux - 2 * ?lpi)))"
@@ -1387,7 +1387,7 @@
also have "\<dots> \<le> cos ((?lx + 2 * ?lpi))"
using cos_monotone_0_pi'[OF lx_0 lx_le_x pi_x]
by (simp only: minus_float.rep_eq real_of_int_minus real_of_one
- minus_one[symmetric] diff_minus mult_minus_left mult_1_left)
+ minus_one[symmetric] mult_minus_left mult_1_left) simp
also have "\<dots> \<le> (ub_cos prec (?lx + 2 * ?lpi))"
using lb_cos[OF lx_0 pi_lx] by simp
finally show ?thesis unfolding u by (simp add: real_of_float_max)
@@ -2164,12 +2164,12 @@
unfolding divide_inverse interpret_floatarith.simps ..
lemma interpret_floatarith_diff: "interpret_floatarith (Add a (Minus b)) vs = (interpret_floatarith a vs) - (interpret_floatarith b vs)"
- unfolding diff_minus interpret_floatarith.simps ..
+ unfolding interpret_floatarith.simps by simp
lemma interpret_floatarith_sin: "interpret_floatarith (Cos (Add (Mult Pi (Num (Float 1 -1))) (Minus a))) vs =
sin (interpret_floatarith a vs)"
unfolding sin_cos_eq interpret_floatarith.simps
- interpret_floatarith_divide interpret_floatarith_diff diff_minus
+ interpret_floatarith_divide interpret_floatarith_diff
by auto
lemma interpret_floatarith_tan:
@@ -3192,7 +3192,7 @@
from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
- by (auto simp add: diff_minus)
+ by auto
from order_less_le_trans[OF _ this, of 0] `0 < ly`
show ?thesis by auto
qed
@@ -3214,7 +3214,7 @@
from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'
have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"
- by (auto simp add: diff_minus)
+ by auto
from order_trans[OF _ this, of 0] `0 \<le> ly`
show ?thesis by auto
qed
--- a/src/HOL/Decision_Procs/Cooper.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Decision_Procs/Cooper.thy Fri Nov 01 18:51:14 2013 +0100
@@ -1400,9 +1400,8 @@
also have "\<dots> = (j dvd (- (c*x - ?e)))"
by (simp only: dvd_minus_iff)
also have "\<dots> = (j dvd (c* (- x)) + ?e)"
- apply (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] diff_minus add_ac minus_add_distrib)
- apply (simp add: algebra_simps)
- done
+ by (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] add_ac minus_add_distrib)
+ (simp add: algebra_simps)
also have "\<dots> = Ifm bbs ((- x)#bs) (Dvd j (CN 0 c e))"
using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] by simp
finally show ?case .
@@ -1413,9 +1412,8 @@
also have "\<dots> = (j dvd (- (c*x - ?e)))"
by (simp only: dvd_minus_iff)
also have "\<dots> = (j dvd (c* (- x)) + ?e)"
- apply (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] diff_minus add_ac minus_add_distrib)
- apply (simp add: algebra_simps)
- done
+ by (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] add_ac minus_add_distrib)
+ (simp add: algebra_simps)
also have "\<dots> = Ifm bbs ((- x)#bs) (Dvd j (CN 0 c e))"
using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] by simp
finally show ?case by simp
--- a/src/HOL/Decision_Procs/MIR.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Decision_Procs/MIR.thy Fri Nov 01 18:51:14 2013 +0100
@@ -1727,7 +1727,7 @@
{assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))"
by (simp add: nb Let_def split_def isint_Floor isint_neg)
have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def)
- also have "\<dots> = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def diff_minus)
+ also have "\<dots> = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def)
finally have ?case using l by simp}
moreover
{assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))"
@@ -1752,13 +1752,13 @@
{assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))"
by (simp add: nb Let_def split_def isint_Floor isint_neg)
have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
- also have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_minus)
+ also have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def)
finally have ?case using l by simp}
moreover
{assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))"
by (simp add: nb Let_def split_def isint_Floor isint_neg)
have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
- also from cn cnz have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac ,arith)
+ also from cn cnz have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def add_ac, arith)
finally have ?case using l by simp}
ultimately show ?case by blast
next
@@ -1777,13 +1777,13 @@
{assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))"
by (simp add: nb Let_def split_def isint_Floor isint_neg)
have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
- also have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_minus)
+ also have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def)
finally have ?case using l by simp}
moreover
{assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))"
by (simp add: nb Let_def split_def isint_Floor isint_neg)
have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
- also from cn cnz have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac, arith)
+ also from cn cnz have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def add_ac, arith)
finally have ?case using l by simp}
ultimately show ?case by blast
next
@@ -1802,13 +1802,13 @@
{assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))"
by (simp add: nb Let_def split_def isint_Floor isint_neg)
have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
- also have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_minus)
+ also have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def)
finally have ?case using l by simp}
moreover
{assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))"
by (simp add: nb Let_def split_def isint_Floor isint_neg)
have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
- also from cn cnz have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac, arith)
+ also from cn cnz have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def add_ac, arith)
finally have ?case using l by simp}
ultimately show ?case by blast
next
@@ -3125,7 +3125,8 @@
hence pid: "c*i + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + ?fe = j1" by auto
hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - ?N i e + real j1"
- by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] algebra_simps)
+ by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff])
+ (simp add: algebra_simps)
with nob have ?case by blast }
ultimately show ?case by blast
next
@@ -3148,11 +3149,12 @@
hence pid: "c*i + 1 + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + 1+ ?fe = j1" by auto
hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) + 1= - ?N i e + real j1"
- by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] algebra_simps real_of_one)
+ by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] real_of_one)
+ (simp add: algebra_simps)
hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1"
by (simp only: algebra_simps)
hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1"
- by (simp only: add_ac diff_minus)
+ by (simp add: algebra_simps minus_one [symmetric] del: minus_one)
with nob have ?case by blast }
ultimately show ?case by blast
next
@@ -3477,10 +3479,7 @@
qed
next
case (3 a b) then show ?case
- apply auto
- apply (erule_tac x = "(aa, aaa, ba)" in ballE) apply simp_all
- apply (erule_tac x = "(ab, ac, baa)" in ballE) apply simp_all
- done
+ by auto
qed (auto simp add: Let_def split_def algebra_simps)
lemma real_in_int_intervals:
@@ -3615,7 +3614,7 @@
by(simp only: myle[of _ "real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"])
hence "\<exists> j\<in> {n .. 0}. 0 \<ge> - (real n *x + ?N s - ?N (Floor s) - real j) \<and> - (real n *x + ?N s - ?N (Floor s) - real (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format])
hence "\<exists> j\<in> {n.. 0}. ?I (?p (p,n,s) j)"
- using pns by (simp add: fp_def nn diff_minus add_ac mult_ac
+ using pns by (simp add: fp_def nn algebra_simps
del: diff_less_0_iff_less diff_le_0_iff_le)
then obtain "j" where j_def: "j\<in> {n .. 0} \<and> ?I (?p (p,n,s) j)" by blast
hence "\<exists>x \<in> {?p (p,n,s) j |j. n\<le> j \<and> j \<le> 0 }. ?I x" by auto
@@ -4832,7 +4831,7 @@
shows "(Ifm bs (E p)) = (\<exists> (i::int). Ifm (real i#bs) (E (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (exsplit p))))" (is "?lhs = ?rhs")
proof-
have "?rhs = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm (x#(real i)#bs) (exsplit p))"
- by (simp add: myless[of _ "1"] myless[of _ "0"] add_ac diff_minus)
+ by (simp add: myless[of _ "1"] myless[of _ "0"] add_ac)
also have "\<dots> = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm ((real i + x) #bs) p)"
by (simp only: exsplit[OF qf] add_ac)
also have "\<dots> = (\<exists> x. Ifm (x#bs) p)"
@@ -5196,7 +5195,7 @@
hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)"
by (auto simp only: subst0_bound0[OF qfmq])
hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))"
- by (auto simp add: simpfm_bound0)
+ by auto
from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp
from Bn jsnb have "\<forall> (b,j) \<in> set ?bjs. numbound0 (Add b (C j))"
by simp
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Fri Nov 01 18:51:14 2013 +0100
@@ -1959,7 +1959,7 @@
by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r = 0"
- by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
+ by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) =0 "
using c d mult_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r =0"
@@ -2041,7 +2041,7 @@
by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r \<noteq> 0"
- by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
+ by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) \<noteq> 0 "
using c d mult_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r \<noteq> 0"
@@ -2106,7 +2106,7 @@
by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r < 0"
- by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
+ by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) < 0"
using dc' dc'' mult_less_cancel_left_disj[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
@@ -2127,7 +2127,7 @@
by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r < 0"
- by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
+ by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) > 0"
@@ -2251,7 +2251,7 @@
by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r <= 0"
- by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
+ by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) <= 0"
using dc' dc'' mult_le_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp
@@ -2272,7 +2272,7 @@
by (simp add: field_simps)
have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r <= 0"
- by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"])
+ by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"])
also have "\<dots> \<longleftrightarrow> (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) >= 0"
@@ -2356,8 +2356,11 @@
lemma msubst_I: assumes lp: "islin p" and nc: "isnpoly c" and nd: "isnpoly d"
shows "Ifm vs (x#bs) (msubst p ((c,t),(d,s))) = Ifm vs (((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /2)#bs) p"
- using lp
-by (induct p rule: islin.induct, auto simp add: tmbound0_I[where b="(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) + - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) / 2" and b'=x and bs = bs and vs=vs] bound0_I[where b="(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) + - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) / 2" and b'=x and bs = bs and vs=vs] msubsteq msubstneq msubstlt[OF nc nd] msubstle[OF nc nd])
+ using lp by (induct p rule: islin.induct)
+ (auto simp add: tmbound0_I
+ [where b = "(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) / 2"
+ and b' = x and bs = bs and vs = vs]
+ msubsteq msubstneq msubstlt [OF nc nd] msubstle [OF nc nd])
lemma msubst_nb: assumes lp: "islin p" and t: "tmbound0 t" and s: "tmbound0 s"
shows "bound0 (msubst p ((c,t),(d,s)))"
@@ -2429,7 +2432,7 @@
with evaldjf_bound0[of ?Up "(simpfm o (msubst (simpfm p)))"]
have "bound0 (evaldjf (simpfm o (msubst (simpfm p))) ?Up)" by blast
with mp_nb pp_nb
- have th1: "bound0 (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up )))" by (simp add: disj_nb)
+ have th1: "bound0 (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up )))" by simp
from decr0_qf[OF th1] have thqf: "qfree (ferrack p)" by (simp add: ferrack_def Let_def)
have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)" by simp
also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> (\<exists>(c, t)\<in>set ?U. \<exists>(d, s)\<in>set ?U. ?I (msubst (simpfm p) ((c, t), d, s)))" using fr_eq_msubst[OF lq, of vs bs x] by simp
@@ -2612,7 +2615,7 @@
lemma msubst2_nb: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" and lp: "islin p" and tnb: "tmbound0 t"
shows "bound0 (msubst2 p c t)"
using lp tnb
-by (simp add: msubst2_def msubstneg_nb msubstpos_nb conj_nb disj_nb lt_nb simpfm_bound0)
+by (simp add: msubst2_def msubstneg_nb msubstpos_nb lt_nb simpfm_bound0)
lemma mult_minus2_left: "-2 * (x::'a::comm_ring_1) = - (2 * x)"
by simp
@@ -2666,8 +2669,8 @@
using H(3) by (auto simp add: msubst2_def lt[OF stupid(1)] lt[OF stupid(2)] zero_less_mult_iff mult_less_0_iff)
from msubst2[OF lp nn nn'(1), of x bs ] H(3) nn'
have "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p"
- apply (simp add: add_divide_distrib mult_minus2_left)
- by (simp add: mult_commute)}
+ by (simp add: add_divide_distrib diff_divide_distrib mult_minus2_left mult_commute)
+ }
moreover
{fix c t d s assume H: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)"
"\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / 2 # bs) p"
@@ -2675,7 +2678,9 @@
hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)" "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
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)))" apply (simp add: add_divide_distrib mult_minus2_left) by (simp add: mult_commute)}
+ 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)
+ }
ultimately show ?thesis by blast
qed
from fr_eq2[OF lp, of vs bs x] show ?thesis
--- a/src/HOL/Decision_Procs/Polynomial_List.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Decision_Procs/Polynomial_List.thy Fri Nov 01 18:51:14 2013 +0100
@@ -419,7 +419,7 @@
lemma (in comm_ring_1) poly_add_minus_mult_eq:
"poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))"
- by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult distrib_left)
+ by (auto simp add: poly_add poly_minus_def fun_eq poly_mult poly_cmult algebra_simps)
subclass (in idom_char_0) comm_ring_1 ..
@@ -445,10 +445,7 @@
lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a,1] \<noteq> poly []"
apply (simp add: fun_eq)
apply (rule_tac x = "minus one a" in exI)
- apply (unfold diff_minus)
- apply (subst add_commute)
- apply (subst add_assoc)
- apply simp
+ apply (simp add: add_commute [of a])
done
lemma (in idom) poly_exp_prime_eq_zero: "poly ([a, 1] %^ n) \<noteq> poly []"
@@ -639,7 +636,7 @@
have "[- a, 1] %^ n divides mulexp n [- a, 1] q"
proof (rule dividesI)
show "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ n *** q)"
- by (induct n) (simp_all add: poly_add poly_cmult poly_mult distrib_left mult_ac)
+ by (induct n) (simp_all add: poly_add poly_cmult poly_mult algebra_simps)
qed
moreover have "\<not> [- a, 1] %^ Suc n divides mulexp n [- a, 1] q"
proof
@@ -873,7 +870,7 @@
proof
assume eq: ?lhs
hence "\<And>x. poly ((c#cs) +++ -- (d#ds)) x = 0"
- by (simp only: poly_minus poly_add algebra_simps) simp
+ by (simp only: poly_minus poly_add algebra_simps) (simp add: algebra_simps)
hence "poly ((c#cs) +++ -- (d#ds)) = poly []" by(simp add: fun_eq_iff)
hence "c = d \<and> list_all (\<lambda>x. x=0) ((cs +++ -- ds))"
unfolding poly_zero by (simp add: poly_minus_def algebra_simps)
--- a/src/HOL/Decision_Procs/Rat_Pair.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Decision_Procs/Rat_Pair.thy Fri Nov 01 18:51:14 2013 +0100
@@ -266,12 +266,13 @@
by (simp add: x y th Nadd_def normNum_def INum_def split_def) }
moreover {
assume z: "a * b' + b * a' \<noteq> 0"
- let ?g = "gcd (a * b' + b * a') (b*b')"
+ let ?g = "gcd (a * b' + b * a') (b * b')"
have gz: "?g \<noteq> 0" using z by simp
have ?thesis using aa' bb' z gz
of_int_div[where ?'a = 'a, OF gz gcd_dvd1_int[where x="a * b' + b * a'" and y="b*b'"]]
of_int_div[where ?'a = 'a, OF gz gcd_dvd2_int[where x="a * b' + b * a'" and y="b*b'"]]
- by (simp add: x y Nadd_def INum_def normNum_def Let_def add_divide_distrib) }
+ by (simp add: x y Nadd_def INum_def normNum_def Let_def) (simp add: field_simps)
+ }
ultimately have ?thesis using aa' bb'
by (simp add: x y Nadd_def INum_def normNum_def Let_def) }
ultimately show ?thesis by blast
--- a/src/HOL/Decision_Procs/mir_tac.ML Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Decision_Procs/mir_tac.ML Fri Nov 01 18:51:14 2013 +0100
@@ -34,7 +34,7 @@
@{thm "divide_zero"},
@{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_right"},
@{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"},
- @{thm "diff_minus"}, @{thm "minus_divide_left"}]
+ @{thm uminus_add_conv_diff [symmetric]}, @{thm "minus_divide_left"}]
val comp_ths = ths @ comp_arith @ @{thms simp_thms};
--- a/src/HOL/Deriv.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Deriv.thy Fri Nov 01 18:51:14 2013 +0100
@@ -98,7 +98,7 @@
lemma FDERIV_diff[simp, FDERIV_intros]:
"(f has_derivative f') F \<Longrightarrow> (g has_derivative g') F \<Longrightarrow> ((\<lambda>x. f x - g x) has_derivative (\<lambda>x. f' x - g' x)) F"
- by (simp only: diff_minus FDERIV_add FDERIV_minus)
+ by (simp only: diff_conv_add_uminus FDERIV_add FDERIV_minus)
abbreviation
-- {* Frechet derivative: D is derivative of function f at x within s *}
@@ -718,13 +718,13 @@
((%x. (f(x)-f(a)) / (x-a)) -- a --> D)"
apply (rule iffI)
apply (drule_tac k="- a" in LIM_offset)
-apply (simp add: diff_minus)
+apply simp
apply (drule_tac k="a" in LIM_offset)
apply (simp add: add_commute)
done
lemma DERIV_iff2: "(DERIV f x :> D) \<longleftrightarrow> (\<lambda>z. (f z - f x) / (z - x)) --x --> D"
- by (simp add: deriv_def diff_minus [symmetric] DERIV_LIM_iff)
+ by (simp add: deriv_def DERIV_LIM_iff)
lemma DERIV_cong_ev: "x = y \<Longrightarrow> eventually (\<lambda>x. f x = g x) (nhds x) \<Longrightarrow> u = v \<Longrightarrow>
DERIV f x :> u \<longleftrightarrow> DERIV g y :> v"
@@ -758,8 +758,7 @@
let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))"
show "\<forall>z. f z - f x = ?g z * (z-x)" by simp
show "isCont ?g x" using der
- by (simp add: isCont_iff DERIV_iff diff_minus
- cong: LIM_equal [rule_format])
+ by (simp add: isCont_iff DERIV_iff cong: LIM_equal [rule_format])
show "?g x = l" by simp
qed
next
@@ -787,7 +786,7 @@
proof -
from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]]
have "\<exists>s > 0. (\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l)"
- by (simp add: diff_minus)
+ by simp
then obtain s
where s: "0 < s"
and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l"
@@ -798,8 +797,7 @@
fix h::real
assume "0 < h" "h < s"
with all [of h] show "f x < f (x+h)"
- proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric]
- split add: split_if_asm)
+ proof (simp add: abs_if pos_less_divide_eq split add: split_if_asm)
assume "~ (f (x+h) - f x) / h < l" and h: "0 < h"
with l
have "0 < (f (x+h) - f x) / h" by arith
@@ -817,7 +815,7 @@
proof -
from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]]
have "\<exists>s > 0. (\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l)"
- by (simp add: diff_minus)
+ by simp
then obtain s
where s: "0 < s"
and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l"
@@ -828,8 +826,7 @@
fix h::real
assume "0 < h" "h < s"
with all [of "-h"] show "f x < f (x-h)"
- proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric]
- split add: split_if_asm)
+ proof (simp add: abs_if pos_less_divide_eq split add: split_if_asm)
assume " - ((f (x-h) - f x) / h) < l" and h: "0 < h"
with l
have "0 < (f (x-h) - f x) / h" by arith
@@ -1131,7 +1128,7 @@
apply (rule linorder_cases [of a b], auto)
apply (drule_tac [!] f = f in MVT)
apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def)
-apply (auto dest: DERIV_unique simp add: ring_distribs diff_minus)
+apply (auto dest: DERIV_unique simp add: ring_distribs)
done
lemma DERIV_const_ratio_const2:
--- a/src/HOL/Divides.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Divides.thy Fri Nov 01 18:51:14 2013 +0100
@@ -439,24 +439,23 @@
text {* Subtraction respects modular equivalence. *}
-lemma mod_diff_left_eq: "(a - b) mod c = (a mod c - b) mod c"
- unfolding diff_minus
- by (intro mod_add_cong mod_minus_cong) simp_all
-
-lemma mod_diff_right_eq: "(a - b) mod c = (a - b mod c) mod c"
- unfolding diff_minus
- by (intro mod_add_cong mod_minus_cong) simp_all
-
-lemma mod_diff_eq: "(a - b) mod c = (a mod c - b mod c) mod c"
- unfolding diff_minus
- by (intro mod_add_cong mod_minus_cong) simp_all
+lemma mod_diff_left_eq:
+ "(a - b) mod c = (a mod c - b) mod c"
+ using mod_add_cong [of a c "a mod c" "- b" "- b"] by simp
+
+lemma mod_diff_right_eq:
+ "(a - b) mod c = (a - b mod c) mod c"
+ using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp
+
+lemma mod_diff_eq:
+ "(a - b) mod c = (a mod c - b mod c) mod c"
+ using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp
lemma mod_diff_cong:
assumes "a mod c = a' mod c"
assumes "b mod c = b' mod c"
shows "(a - b) mod c = (a' - b') mod c"
- unfolding diff_minus using assms
- by (intro mod_add_cong mod_minus_cong)
+ using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"] by simp
lemma dvd_neg_div: "y dvd x \<Longrightarrow> -x div y = - (x div y)"
apply (case_tac "y = 0") apply simp
@@ -502,10 +501,7 @@
lemma minus_mod_self1 [simp]:
"(b - a) mod b = - a mod b"
-proof -
- have "b - a = - a + b" by (simp add: diff_minus add.commute)
- then show ?thesis by simp
-qed
+ using mod_add_self2 [of "- a" b] by simp
end
@@ -1749,7 +1745,7 @@
val div_mod_eqs = map mk_meta_eq [@{thm div_mod_equality}, @{thm div_mod_equality2}];
val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac
- (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac}))
+ (@{thm diff_conv_add_uminus} :: @{thms add_0s} @ @{thms add_ac}))
)
*}
--- a/src/HOL/Fields.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Fields.thy Fri Nov 01 18:51:14 2013 +0100
@@ -156,7 +156,7 @@
by (simp add: divide_inverse)
lemma diff_divide_distrib: "(a - b) / c = a / c - b / c"
- by (simp add: diff_minus add_divide_distrib)
+ using add_divide_distrib [of a "- b" c] by simp
lemma nonzero_eq_divide_eq [field_simps]: "c \<noteq> 0 \<Longrightarrow> a = b / c \<longleftrightarrow> a * c = b"
proof -
@@ -845,7 +845,7 @@
fix x y :: 'a
from less_add_one show "\<exists>y. x < y" ..
from less_add_one have "x + (- 1) < (x + 1) + (- 1)" by (rule add_strict_right_mono)
- then have "x - 1 < x + 1 - 1" by (simp only: diff_minus [symmetric])
+ then have "x - 1 < x + 1 - 1" by simp
then have "x - 1 < x" by (simp add: algebra_simps)
then show "\<exists>y. y < x" ..
show "x < y \<Longrightarrow> \<exists>z>x. z < y" by (blast intro!: less_half_sum gt_half_sum)
--- a/src/HOL/Groups.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Groups.thy Fri Nov 01 18:51:14 2013 +0100
@@ -321,9 +321,13 @@
class group_add = minus + uminus + monoid_add +
assumes left_minus [simp]: "- a + a = 0"
- assumes diff_minus: "a - b = a + (- b)"
+ assumes add_uminus_conv_diff [simp]: "a + (- b) = a - b"
begin
+lemma diff_conv_add_uminus:
+ "a - b = a + (- b)"
+ by simp
+
lemma minus_unique:
assumes "a + b = 0" shows "- a = b"
proof -
@@ -332,8 +336,6 @@
finally show ?thesis .
qed
-lemmas equals_zero_I = minus_unique (* legacy name *)
-
lemma minus_zero [simp]: "- 0 = 0"
proof -
have "0 + 0 = 0" by (rule add_0_right)
@@ -346,13 +348,17 @@
thus "- (- a) = a" by (rule minus_unique)
qed
-lemma right_minus [simp]: "a + - a = 0"
+lemma right_minus: "a + - a = 0"
proof -
have "a + - a = - (- a) + - a" by simp
also have "\<dots> = 0" by (rule left_minus)
finally show ?thesis .
qed
+lemma diff_self [simp]:
+ "a - a = 0"
+ using right_minus [of a] by simp
+
subclass cancel_semigroup_add
proof
fix a b c :: 'a
@@ -367,41 +373,57 @@
then show "b = c" unfolding add_assoc by simp
qed
-lemma minus_add_cancel: "- a + (a + b) = b"
-by (simp add: add_assoc [symmetric])
+lemma minus_add_cancel [simp]:
+ "- a + (a + b) = b"
+ by (simp add: add_assoc [symmetric])
+
+lemma add_minus_cancel [simp]:
+ "a + (- a + b) = b"
+ by (simp add: add_assoc [symmetric])
-lemma add_minus_cancel: "a + (- a + b) = b"
-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
-lemma minus_add: "- (a + b) = - b + - a"
+lemma add_diff_cancel [simp]:
+ "a + b - b = a"
+ 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 add: add_assoc add_minus_cancel)
- thus "- (a + b) = - b + - a"
+ by (simp only: add_assoc add_minus_cancel) simp
+ then show "- (a + b) = - b + - a"
by (rule minus_unique)
qed
-lemma right_minus_eq: "a - b = 0 \<longleftrightarrow> a = b"
+lemma right_minus_eq [simp]:
+ "a - b = 0 \<longleftrightarrow> a = b"
proof
assume "a - b = 0"
- have "a = (a - b) + b" by (simp add:diff_minus add_assoc)
+ have "a = (a - b) + b" by (simp add: add_assoc)
also have "\<dots> = b" using `a - b = 0` by simp
finally show "a = b" .
next
- assume "a = b" thus "a - b = 0" by (simp add: diff_minus)
+ assume "a = b" thus "a - b = 0" by simp
qed
-lemma diff_self [simp]: "a - a = 0"
-by (simp add: diff_minus)
+lemma eq_iff_diff_eq_0:
+ "a = b \<longleftrightarrow> a - b = 0"
+ by (fact right_minus_eq [symmetric])
-lemma diff_0 [simp]: "0 - a = - a"
-by (simp add: diff_minus)
+lemma diff_0 [simp]:
+ "0 - a = - a"
+ by (simp only: diff_conv_add_uminus add_0_left)
-lemma diff_0_right [simp]: "a - 0 = a"
-by (simp add: diff_minus)
+lemma diff_0_right [simp]:
+ "a - 0 = a"
+ by (simp only: diff_conv_add_uminus minus_zero add_0_right)
-lemma diff_minus_eq_add [simp]: "a - - b = a + b"
-by (simp add: diff_minus)
+lemma diff_minus_eq_add [simp]:
+ "a - - b = a + b"
+ by (simp only: diff_conv_add_uminus minus_minus)
lemma neg_equal_iff_equal [simp]:
"- a = - b \<longleftrightarrow> a = b"
@@ -416,11 +438,11 @@
lemma neg_equal_0_iff_equal [simp]:
"- a = 0 \<longleftrightarrow> a = 0"
-by (subst neg_equal_iff_equal [symmetric], simp)
+ by (subst neg_equal_iff_equal [symmetric]) simp
lemma neg_0_equal_iff_equal [simp]:
"0 = - a \<longleftrightarrow> 0 = a"
-by (subst neg_equal_iff_equal [symmetric], simp)
+ by (subst neg_equal_iff_equal [symmetric]) simp
text{*The next two equations can make the simplifier loop!*}
@@ -438,15 +460,8 @@
thus ?thesis by (simp add: eq_commute)
qed
-lemma diff_add_cancel: "a - b + b = a"
-by (simp add: diff_minus add_assoc)
-
-lemma add_diff_cancel: "a + b - b = a"
-by (simp add: diff_minus add_assoc)
-
-declare diff_minus[symmetric, algebra_simps, field_simps]
-
-lemma eq_neg_iff_add_eq_0: "a = - b \<longleftrightarrow> a + b = 0"
+lemma eq_neg_iff_add_eq_0:
+ "a = - b \<longleftrightarrow> a + b = 0"
proof
assume "a = - b" then show "a + b = 0" by simp
next
@@ -456,72 +471,88 @@
ultimately show "a = - b" by simp
qed
-lemma add_eq_0_iff: "x + y = 0 \<longleftrightarrow> y = - x"
- unfolding eq_neg_iff_add_eq_0 [symmetric]
- by (rule equation_minus_iff)
+lemma add_eq_0_iff2:
+ "a + b = 0 \<longleftrightarrow> a = - b"
+ by (fact eq_neg_iff_add_eq_0 [symmetric])
+
+lemma neg_eq_iff_add_eq_0:
+ "- a = b \<longleftrightarrow> a + b = 0"
+ by (auto simp add: add_eq_0_iff2)
-lemma minus_diff_eq [simp]: "- (a - b) = b - a"
- by (simp add: diff_minus minus_add)
+lemma add_eq_0_iff:
+ "a + b = 0 \<longleftrightarrow> b = - a"
+ by (auto simp add: neg_eq_iff_add_eq_0 [symmetric])
-lemma add_diff_eq[algebra_simps, field_simps]: "a + (b - c) = (a + b) - c"
- by (simp add: diff_minus add_assoc)
+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
-lemma diff_eq_eq[algebra_simps, field_simps]: "a - b = c \<longleftrightarrow> a = c + b"
- by (auto simp add: diff_minus add_assoc)
+lemma add_diff_eq [algebra_simps, field_simps]:
+ "a + (b - c) = (a + b) - c"
+ by (simp only: diff_conv_add_uminus add_assoc)
-lemma eq_diff_eq[algebra_simps, field_simps]: "a = c - b \<longleftrightarrow> a + b = c"
- by (auto simp add: diff_minus 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)
-lemma diff_diff_eq2[algebra_simps, field_simps]: "a - (b - c) = (a + c) - b"
- by (simp add: diff_minus minus_add add_assoc)
+lemma diff_eq_eq [algebra_simps, field_simps]:
+ "a - b = c \<longleftrightarrow> a = c + b"
+ by auto
-lemma eq_iff_diff_eq_0: "a = b \<longleftrightarrow> a - b = 0"
- by (fact right_minus_eq [symmetric])
+lemma eq_diff_eq [algebra_simps, field_simps]:
+ "a = c - b \<longleftrightarrow> a + b = c"
+ by auto
+
+lemma diff_diff_eq2 [algebra_simps, field_simps]:
+ "a - (b - c) = (a + c) - b"
+ by (simp only: diff_conv_add_uminus add_assoc) simp
lemma diff_eq_diff_eq:
"a - b = c - d \<Longrightarrow> a = b \<longleftrightarrow> c = d"
- by (simp add: eq_iff_diff_eq_0 [of a b] eq_iff_diff_eq_0 [of c d])
+ by (simp only: eq_iff_diff_eq_0 [of a b] eq_iff_diff_eq_0 [of c d])
end
class ab_group_add = minus + uminus + comm_monoid_add +
assumes ab_left_minus: "- a + a = 0"
- assumes ab_diff_minus: "a - b = a + (- b)"
+ assumes ab_add_uminus_conv_diff: "a - b = a + (- b)"
begin
subclass group_add
- proof qed (simp_all add: ab_left_minus ab_diff_minus)
+ proof qed (simp_all add: ab_left_minus ab_add_uminus_conv_diff)
subclass cancel_comm_monoid_add
proof
fix a b c :: 'a
assume "a + b = a + c"
then have "- a + a + b = - a + a + c"
- unfolding add_assoc by simp
+ by (simp only: add_assoc)
then show "b = c" by simp
qed
-lemma uminus_add_conv_diff[algebra_simps, field_simps]:
+lemma uminus_add_conv_diff [simp]:
"- a + b = b - a"
-by (simp add:diff_minus add_commute)
+ by (simp add: add_commute)
lemma minus_add_distrib [simp]:
"- (a + b) = - a + - b"
-by (rule minus_unique) (simp add: add_ac)
+ by (simp add: algebra_simps)
-lemma diff_add_eq[algebra_simps, field_simps]: "(a - b) + c = (a + c) - b"
-by (simp add: diff_minus add_ac)
-
-lemma diff_diff_eq[algebra_simps, field_simps]: "(a - b) - c = a - (b + c)"
-by (simp add: diff_minus add_ac)
+lemma diff_add_eq [algebra_simps, field_simps]:
+ "(a - b) + c = (a + c) - b"
+ by (simp add: algebra_simps)
-(* FIXME: duplicates right_minus_eq from class group_add *)
-(* but only this one is declared as a simp rule. *)
-lemma diff_eq_0_iff_eq [simp]: "a - b = 0 \<longleftrightarrow> a = b"
- by (rule right_minus_eq)
+lemma diff_diff_eq [algebra_simps, field_simps]:
+ "(a - b) - c = a - (b + c)"
+ by (simp add: algebra_simps)
-lemma add_diff_cancel_left: "(c + a) - (c + b) = a - b"
- by (simp add: diff_minus add_ac)
+lemma diff_add_eq_diff_diff:
+ "a - (b + c) = a - b - c"
+ using diff_add_eq_diff_diff_swap [of a c b] by (simp add: add.commute)
+
+lemma add_diff_cancel_left [simp]:
+ "(c + a) - (c + b) = a - b"
+ by (simp add: algebra_simps)
end
@@ -622,19 +653,19 @@
lemma add_less_cancel_left [simp]:
"c + a < c + b \<longleftrightarrow> a < b"
-by (blast intro: add_less_imp_less_left add_strict_left_mono)
+ by (blast intro: add_less_imp_less_left add_strict_left_mono)
lemma add_less_cancel_right [simp]:
"a + c < b + c \<longleftrightarrow> a < b"
-by (blast intro: add_less_imp_less_right add_strict_right_mono)
+ by (blast intro: add_less_imp_less_right add_strict_right_mono)
lemma add_le_cancel_left [simp]:
"c + a \<le> c + b \<longleftrightarrow> a \<le> b"
-by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono)
+ by (auto, drule add_le_imp_le_left, simp_all add: add_left_mono)
lemma add_le_cancel_right [simp]:
"a + c \<le> b + c \<longleftrightarrow> a \<le> 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 \<le> b + c \<Longrightarrow> a \<le> b"
@@ -806,6 +837,22 @@
then show "x + y = 0" by simp
qed
+lemma add_increasing:
+ "0 \<le> a \<Longrightarrow> b \<le> c \<Longrightarrow> b \<le> a + c"
+ by (insert add_mono [of 0 a b c], simp)
+
+lemma add_increasing2:
+ "0 \<le> c \<Longrightarrow> b \<le> a \<Longrightarrow> b \<le> a + c"
+ by (simp add: add_increasing add_commute [of a])
+
+lemma add_strict_increasing:
+ "0 < a \<Longrightarrow> b \<le> c \<Longrightarrow> b < a + c"
+ by (insert add_less_le_mono [of 0 a b c], simp)
+
+lemma add_strict_increasing2:
+ "0 \<le> a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"
+ by (insert add_le_less_mono [of 0 a b c], simp)
+
end
class ordered_ab_group_add =
@@ -825,21 +872,53 @@
subclass ordered_comm_monoid_add ..
+lemma add_less_same_cancel1 [simp]:
+ "b + a < b \<longleftrightarrow> a < 0"
+ using add_less_cancel_left [of _ _ 0] by simp
+
+lemma add_less_same_cancel2 [simp]:
+ "a + b < b \<longleftrightarrow> a < 0"
+ using add_less_cancel_right [of _ _ 0] by simp
+
+lemma less_add_same_cancel1 [simp]:
+ "a < a + b \<longleftrightarrow> 0 < b"
+ using add_less_cancel_left [of _ 0] by simp
+
+lemma less_add_same_cancel2 [simp]:
+ "a < b + a \<longleftrightarrow> 0 < b"
+ using add_less_cancel_right [of 0] by simp
+
+lemma add_le_same_cancel1 [simp]:
+ "b + a \<le> b \<longleftrightarrow> a \<le> 0"
+ using add_le_cancel_left [of _ _ 0] by simp
+
+lemma add_le_same_cancel2 [simp]:
+ "a + b \<le> b \<longleftrightarrow> a \<le> 0"
+ using add_le_cancel_right [of _ _ 0] by simp
+
+lemma le_add_same_cancel1 [simp]:
+ "a \<le> a + b \<longleftrightarrow> 0 \<le> b"
+ using add_le_cancel_left [of _ 0] by simp
+
+lemma le_add_same_cancel2 [simp]:
+ "a \<le> b + a \<longleftrightarrow> 0 \<le> b"
+ using add_le_cancel_right [of 0] by simp
+
lemma max_diff_distrib_left:
shows "max x y - z = max (x - z) (y - z)"
-by (simp add: diff_minus, rule max_add_distrib_left)
+ using max_add_distrib_left [of x y "- z"] by simp
lemma min_diff_distrib_left:
shows "min x y - z = min (x - z) (y - z)"
-by (simp add: diff_minus, rule min_add_distrib_left)
+ using min_add_distrib_left [of x y "- z"] by simp
lemma le_imp_neg_le:
assumes "a \<le> b" shows "-b \<le> -a"
proof -
have "-a+a \<le> -a+b" using `a \<le> b` by (rule add_left_mono)
- hence "0 \<le> -a+b" by simp
- hence "0 + (-b) \<le> (-a + b) + (-b)" by (rule add_right_mono)
- thus ?thesis by (simp add: add_assoc)
+ then have "0 \<le> -a+b" by simp
+ then have "0 + (-b) \<le> (-a + b) + (-b)" by (rule add_right_mono)
+ then show ?thesis by (simp add: algebra_simps)
qed
lemma neg_le_iff_le [simp]: "- b \<le> - a \<longleftrightarrow> a \<le> b"
@@ -899,30 +978,32 @@
lemma diff_less_0_iff_less [simp]:
"a - b < 0 \<longleftrightarrow> a < b"
proof -
- have "a - b < 0 \<longleftrightarrow> a + (- b) < b + (- b)" by (simp add: diff_minus)
+ have "a - b < 0 \<longleftrightarrow> a + (- b) < b + (- b)" by simp
also have "... \<longleftrightarrow> a < b" by (simp only: add_less_cancel_right)
finally show ?thesis .
qed
lemmas less_iff_diff_less_0 = diff_less_0_iff_less [symmetric]
-lemma diff_less_eq[algebra_simps, field_simps]: "a - b < c \<longleftrightarrow> a < c + b"
+lemma diff_less_eq [algebra_simps, field_simps]:
+ "a - b < c \<longleftrightarrow> a < c + b"
apply (subst less_iff_diff_less_0 [of a])
apply (rule less_iff_diff_less_0 [of _ c, THEN ssubst])
-apply (simp add: diff_minus add_ac)
+apply (simp add: algebra_simps)
done
-lemma less_diff_eq[algebra_simps, field_simps]: "a < c - b \<longleftrightarrow> a + b < c"
+lemma less_diff_eq[algebra_simps, field_simps]:
+ "a < c - b \<longleftrightarrow> a + b < c"
apply (subst less_iff_diff_less_0 [of "a + b"])
apply (subst less_iff_diff_less_0 [of a])
-apply (simp add: diff_minus add_ac)
+apply (simp add: algebra_simps)
done
lemma diff_le_eq[algebra_simps, field_simps]: "a - b \<le> c \<longleftrightarrow> a \<le> c + b"
-by (auto simp add: le_less diff_less_eq diff_add_cancel add_diff_cancel)
+by (auto simp add: le_less diff_less_eq )
lemma le_diff_eq[algebra_simps, field_simps]: "a \<le> c - b \<longleftrightarrow> a + b \<le> c"
-by (auto simp add: le_less less_diff_eq diff_add_cancel add_diff_cancel)
+by (auto simp add: le_less less_diff_eq)
lemma diff_le_0_iff_le [simp]:
"a - b \<le> 0 \<longleftrightarrow> a \<le> b"
@@ -1118,21 +1199,11 @@
lemma le_minus_self_iff:
"a \<le> - a \<longleftrightarrow> a \<le> 0"
-proof -
- from add_le_cancel_left [of "- a" "a + a" 0]
- have "a \<le> - a \<longleftrightarrow> a + a \<le> 0"
- by (simp add: add_assoc [symmetric])
- thus ?thesis by simp
-qed
+ by (fact less_eq_neg_nonpos)
lemma minus_le_self_iff:
"- a \<le> a \<longleftrightarrow> 0 \<le> a"
-proof -
- from add_le_cancel_left [of "- a" 0 "a + a"]
- have "- a \<le> a \<longleftrightarrow> 0 \<le> a + a"
- by (simp add: add_assoc [symmetric])
- thus ?thesis by simp
-qed
+ by (fact neg_less_eq_nonneg)
lemma minus_max_eq_min:
"- max x y = min (-x) (-y)"
@@ -1144,27 +1215,6 @@
end
-context ordered_comm_monoid_add
-begin
-
-lemma add_increasing:
- "0 \<le> a \<Longrightarrow> b \<le> c \<Longrightarrow> b \<le> a + c"
- by (insert add_mono [of 0 a b c], simp)
-
-lemma add_increasing2:
- "0 \<le> c \<Longrightarrow> b \<le> a \<Longrightarrow> b \<le> a + c"
- by (simp add: add_increasing add_commute [of a])
-
-lemma add_strict_increasing:
- "0 < a \<Longrightarrow> b \<le> c \<Longrightarrow> b < a + c"
- by (insert add_less_le_mono [of 0 a b c], simp)
-
-lemma add_strict_increasing2:
- "0 \<le> a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"
- by (insert add_le_less_mono [of 0 a b c], simp)
-
-end
-
class abs =
fixes abs :: "'a \<Rightarrow> 'a"
begin
@@ -1299,7 +1349,7 @@
lemma abs_triangle_ineq2: "\<bar>a\<bar> - \<bar>b\<bar> \<le> \<bar>a - b\<bar>"
proof -
have "\<bar>a\<bar> = \<bar>b + (a - b)\<bar>"
- by (simp add: algebra_simps add_diff_cancel)
+ by (simp add: algebra_simps)
then have "\<bar>a\<bar> \<le> \<bar>b\<bar> + \<bar>a - b\<bar>"
by (simp add: abs_triangle_ineq)
then show ?thesis
@@ -1314,14 +1364,14 @@
lemma abs_triangle_ineq4: "\<bar>a - b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
proof -
- have "\<bar>a - b\<bar> = \<bar>a + - b\<bar>" by (subst diff_minus, rule refl)
+ have "\<bar>a - b\<bar> = \<bar>a + - b\<bar>" by (simp add: algebra_simps)
also have "... \<le> \<bar>a\<bar> + \<bar>- b\<bar>" by (rule abs_triangle_ineq)
finally show ?thesis by simp
qed
lemma abs_diff_triangle_ineq: "\<bar>a + b - (c + d)\<bar> \<le> \<bar>a - c\<bar> + \<bar>b - d\<bar>"
proof -
- have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: diff_minus add_ac)
+ have "\<bar>a + b - (c+d)\<bar> = \<bar>(a-c) + (b-d)\<bar>" by (simp add: algebra_simps)
also have "... \<le> \<bar>a-c\<bar> + \<bar>b-d\<bar>" by (rule abs_triangle_ineq)
finally show ?thesis .
qed
@@ -1362,10 +1412,5 @@
code_identifier
code_module Groups \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
-
-text {* Legacy *}
-
-lemmas diff_def = diff_minus
-
end
--- a/src/HOL/Hahn_Banach/Vector_Space.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Hahn_Banach/Vector_Space.thy Fri Nov 01 18:51:14 2013 +0100
@@ -112,7 +112,7 @@
proof -
assume x: "x \<in> V"
have " (a - b) \<cdot> x = (a + - b) \<cdot> x"
- by (simp add: diff_minus)
+ by simp
also from x have "\<dots> = a \<cdot> x + (- b) \<cdot> x"
by (rule add_mult_distrib2)
also from x have "\<dots> = a \<cdot> x + - (b \<cdot> x)"
--- a/src/HOL/Int.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Int.thy Fri Nov 01 18:51:14 2013 +0100
@@ -220,7 +220,7 @@
by (transfer fixing: uminus) clarsimp
lemma of_int_diff [simp]: "of_int (w - z) = of_int w - of_int z"
-by (simp add: diff_minus Groups.diff_minus)
+ using of_int_add [of w "- z"] by simp
lemma of_int_mult [simp]: "of_int (w*z) = of_int w * of_int z"
by (transfer fixing: times) (clarsimp simp add: algebra_simps of_nat_mult)
@@ -749,7 +749,7 @@
lemmas int_arith_rules =
neg_le_iff_le numeral_One
- minus_zero diff_minus left_minus right_minus
+ minus_zero left_minus right_minus
mult_zero_left mult_zero_right mult_1_left mult_1_right
mult_minus_left mult_minus_right
minus_add_distrib minus_minus mult_assoc
@@ -793,7 +793,6 @@
subsection{*The functions @{term nat} and @{term int}*}
text{*Simplify the term @{term "w + - z"}*}
-lemmas diff_int_def_symmetric = diff_def [where 'a=int, symmetric, simp]
lemma one_less_nat_eq [simp]: "(Suc 0 < nat z) = (1 < z)"
apply (insert zless_nat_conj [of 1 z])
@@ -1510,10 +1509,13 @@
"sub (Num.Bit1 m) (Num.Bit1 n) = dup (sub m n)"
"sub (Num.Bit1 m) (Num.Bit0 n) = dup (sub m n) + 1"
"sub (Num.Bit0 m) (Num.Bit1 n) = dup (sub m n) - 1"
- unfolding sub_def dup_def numeral.simps Pos_def Neg_def
- neg_numeral_def numeral_BitM
- by (simp_all only: algebra_simps)
-
+ apply (simp_all only: sub_def dup_def numeral.simps Pos_def Neg_def
+ neg_numeral_def numeral_BitM)
+ apply (simp_all only: algebra_simps minus_diff_eq)
+ apply (simp_all only: add.commute [of _ "- (numeral n + numeral n)"])
+ apply (simp_all only: minus_add add.assoc left_minus)
+ apply (simp_all only: algebra_simps right_minus)
+ done
text {* Implementations *}
--- a/src/HOL/Library/BigO.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Library/BigO.thy Fri Nov 01 18:51:14 2013 +0100
@@ -215,7 +215,7 @@
f : lb +o O(g)"
apply (rule set_minus_imp_plus)
apply (rule bigo_bounded)
- apply (auto simp add: diff_minus fun_Compl_def func_plus)
+ apply (auto simp add: fun_Compl_def func_plus)
apply (drule_tac x = x in spec)+
apply force
apply (drule_tac x = x in spec)+
@@ -390,7 +390,7 @@
apply (rule set_minus_imp_plus)
apply (drule set_plus_imp_minus)
apply (drule bigo_minus)
- apply (simp add: diff_minus)
+ apply simp
done
lemma bigo_minus3: "O(-f) = O(f)"
@@ -446,7 +446,7 @@
apply (rule bigo_minus)
apply (subst set_minus_plus)
apply assumption
- apply (simp add: diff_minus add_ac)
+ apply (simp add: add_ac)
done
lemma bigo_add_commute: "(f : g +o O(h)) = (g : f +o O(h))"
@@ -545,10 +545,9 @@
lemma bigo_compose2: "f =o g +o O(h) ==> (%x. f(k x)) =o (%x. g(k x)) +o
O(%x. h(k x))"
- apply (simp only: set_minus_plus [symmetric] diff_minus fun_Compl_def
- func_plus)
- apply (erule bigo_compose1)
-done
+ apply (simp only: set_minus_plus [symmetric] fun_Compl_def func_plus)
+ apply (drule bigo_compose1) apply (simp add: fun_diff_def)
+ done
subsection {* Setsum *}
@@ -779,7 +778,7 @@
apply (subst abs_of_nonneg)
apply (drule_tac x = x in spec) back
apply (simp add: algebra_simps)
- apply (subst diff_minus)+
+ apply (subst diff_conv_add_uminus)+
apply (rule add_right_mono)
apply (erule spec)
apply (rule order_trans)
@@ -803,7 +802,7 @@
apply (subst abs_of_nonneg)
apply (drule_tac x = x in spec) back
apply (simp add: algebra_simps)
- apply (subst diff_minus)+
+ apply (subst diff_conv_add_uminus)+
apply (rule add_left_mono)
apply (rule le_imp_neg_le)
apply (erule spec)
--- a/src/HOL/Library/Convex.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Library/Convex.thy Fri Nov 01 18:51:14 2013 +0100
@@ -362,7 +362,7 @@
shows "convex {x - y| x y. x \<in> s \<and> y \<in> t}"
proof -
have "{x - y| x y. x \<in> s \<and> y \<in> t} = {x + y |x y. x \<in> s \<and> y \<in> uminus ` t}"
- unfolding diff_def by auto
+ by (auto simp add: diff_conv_add_uminus simp del: add_uminus_conv_diff)
then show ?thesis
using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto
qed
--- a/src/HOL/Library/Float.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Library/Float.thy Fri Nov 01 18:51:14 2013 +0100
@@ -88,7 +88,7 @@
done
lemma minus_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x - y \<in> float"
- unfolding ab_diff_minus by (intro uminus_float plus_float)
+ using plus_float [of x "- y"] by simp
lemma abs_float[simp]: "x \<in> float \<Longrightarrow> abs x \<in> float"
by (cases x rule: linorder_cases[of 0]) auto
@@ -960,7 +960,7 @@
also have "... < (1 / 2) * 2 powr real (rat_precision n (int x) (int y))"
apply (rule mult_strict_right_mono) by (insert assms) auto
also have "\<dots> = 2 powr real (rat_precision n (int x) (int y) - 1)"
- by (simp add: powr_add diff_def powr_neg_numeral)
+ using powr_add [of 2 _ "- 1", simplified add_uminus_conv_diff] by (simp add: powr_neg_numeral)
also have "\<dots> = 2 ^ nat (rat_precision n (int x) (int y) - 1)"
using rat_precision_pos[of x y n] assms by (simp add: powr_realpow[symmetric])
also have "\<dots> \<le> 2 ^ nat (rat_precision n (int x) (int y)) - 1"
--- a/src/HOL/Library/Formal_Power_Series.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Library/Formal_Power_Series.thy Fri Nov 01 18:51:14 2013 +0100
@@ -231,7 +231,7 @@
proof
fix a b :: "'a fps"
show "- a + a = 0" by (simp add: fps_ext)
- show "a - b = a + - b" by (simp add: fps_ext diff_minus)
+ show "a + - b = a - b" by (simp add: fps_ext)
qed
instance fps :: (ab_group_add) ab_group_add
@@ -348,10 +348,10 @@
instance fps :: (ring) ring ..
instance fps :: (ring_1) ring_1
- by (intro_classes, auto simp add: diff_minus distrib_right)
+ by (intro_classes, auto simp add: distrib_right)
instance fps :: (comm_ring_1) comm_ring_1
- by (intro_classes, auto simp add: diff_minus distrib_right)
+ by (intro_classes, auto simp add: distrib_right)
instance fps :: (ring_no_zero_divisors) ring_no_zero_divisors
proof
@@ -888,7 +888,7 @@
using fps_deriv_linear[of 1 f 1 g] by simp
lemma fps_deriv_sub[simp]: "fps_deriv ((f:: ('a::comm_ring_1) fps) - g) = fps_deriv f - fps_deriv g"
- unfolding diff_minus by simp
+ using fps_deriv_add [of f "- g"] by simp
lemma fps_deriv_const[simp]: "fps_deriv (fps_const c) = 0"
by (simp add: fps_ext fps_deriv_def fps_const_def)
@@ -978,7 +978,7 @@
lemma fps_nth_deriv_sub[simp]:
"fps_nth_deriv n ((f:: ('a::comm_ring_1) fps) - g) = fps_nth_deriv n f - fps_nth_deriv n g"
- unfolding diff_minus fps_nth_deriv_add by simp
+ using fps_nth_deriv_add [of n f "- g"] by simp
lemma fps_nth_deriv_0[simp]: "fps_nth_deriv n 0 = 0"
by (induct n) simp_all
@@ -2634,7 +2634,7 @@
by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_negf[symmetric])
lemma fps_compose_sub_distrib: "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
- unfolding diff_minus fps_compose_uminus fps_compose_add_distrib ..
+ using fps_compose_add_distrib [of a "- b" c] by (simp add: fps_compose_uminus)
lemma X_fps_compose: "X oo a = Abs_fps (\<lambda>n. if n = 0 then (0::'a::comm_ring_1) else a$n)"
by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum_delta)
--- a/src/HOL/Library/Fraction_Field.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Library/Fraction_Field.thy Fri Nov 01 18:51:14 2013 +0100
@@ -132,7 +132,7 @@
lemma diff_fract [simp]:
assumes "b \<noteq> 0" and "d \<noteq> 0"
shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
- using assms by (simp add: diff_fract_def diff_minus)
+ using assms by (simp add: diff_fract_def)
definition mult_fract_def:
"q * r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
--- a/src/HOL/Library/Function_Algebras.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Library/Function_Algebras.thy Fri Nov 01 18:51:14 2013 +0100
@@ -83,10 +83,10 @@
instance "fun" :: (type, group_add) group_add
by default
- (simp_all add: fun_eq_iff diff_minus)
+ (simp_all add: fun_eq_iff)
instance "fun" :: (type, ab_group_add) ab_group_add
- by default (simp_all add: diff_minus)
+ by default simp_all
text {* Multiplicative structures *}
--- a/src/HOL/Library/Fundamental_Theorem_Algebra.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Library/Fundamental_Theorem_Algebra.thy Fri Nov 01 18:51:14 2013 +0100
@@ -224,12 +224,12 @@
from unimodular_reduce_norm[OF th0] o
have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1", rule_tac x="1" in exI, simp)
- apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp add: diff_minus)
+ apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1", rule_tac x="-1" in exI, simp del: minus_one add: minus_one [symmetric])
apply (cases "cmod (complex_of_real (cmod b) / b + ii) < 1")
apply (cases "even m", rule_tac x="ii" in exI, simp add: m power_mult)
apply (rule_tac x="- ii" in exI, simp add: m power_mult)
- apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult diff_minus)
- apply (rule_tac x="ii" in exI, simp add: m power_mult diff_minus)
+ apply (cases "even m", rule_tac x="- ii" in exI, simp add: m power_mult)
+ apply (rule_tac x="ii" in exI, simp add: m power_mult)
done
then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1" by blast
let ?w = "v / complex_of_real (root n (cmod b))"
@@ -954,7 +954,7 @@
lemma mpoly_sub_conv:
"poly p (x::complex) - poly q x \<equiv> poly p x + -1 * poly q x"
- by (simp add: diff_minus)
+ by simp
lemma poly_pad_rule: "poly p x = 0 ==> poly (pCons 0 p) x = (0::complex)" by simp
--- a/src/HOL/Library/Inner_Product.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Library/Inner_Product.thy Fri Nov 01 18:51:14 2013 +0100
@@ -41,7 +41,7 @@
using inner_add_left [of x "- x" y] by simp
lemma inner_diff_left: "inner (x - y) z = inner x z - inner y z"
- by (simp add: diff_minus inner_add_left)
+ using inner_add_left [of x "- y" z] by simp
lemma inner_setsum_left: "inner (\<Sum>x\<in>A. f x) y = (\<Sum>x\<in>A. inner (f x) y)"
by (cases "finite A", induct set: finite, simp_all add: inner_add_left)
--- a/src/HOL/Library/Lattice_Algebras.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Library/Lattice_Algebras.thy Fri Nov 01 18:51:14 2013 +0100
@@ -13,9 +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)
- apply rule
- apply (rule add_le_imp_le_left[of "a"], simp only: add_assoc[symmetric], simp)+
+ 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)"
@@ -33,11 +31,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 rule
+ 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 le_supI)
- apply (simp_all)
done
lemma add_sup_distrib_right: "sup a b + c = sup (a+c) (b+c)"
@@ -87,9 +84,15 @@
lemma neg_inf_eq_sup: "- inf a b = sup (-a) (-b)"
by (simp add: inf_eq_neg_sup)
+lemma diff_inf_eq_sup: "a - inf b c = a + sup (- b) (- c)"
+ using neg_inf_eq_sup [of b c, symmetric] by simp
+
lemma neg_sup_eq_inf: "- sup a b = inf (-a) (-b)"
by (simp add: sup_eq_neg_inf)
+lemma diff_sup_eq_inf: "a - sup b c = a + inf (- b) (- c)"
+ using neg_sup_eq_inf [of b c, symmetric] by simp
+
lemma add_eq_inf_sup: "a + b = sup a b + inf a b"
proof -
have "0 = - inf 0 (a-b) + inf (a-b) 0"
@@ -97,8 +100,8 @@
hence "0 = sup 0 (b-a) + inf (a-b) 0"
by (simp add: inf_eq_neg_sup)
hence "0 = (-a + sup a b) + (inf a b + (-b))"
- by (simp add: add_sup_distrib_left add_inf_distrib_right) (simp add: algebra_simps)
- thus ?thesis by (simp add: algebra_simps)
+ by (simp only: add_sup_distrib_left add_inf_distrib_right) simp
+ then show ?thesis by (simp add: algebra_simps)
qed
@@ -251,7 +254,7 @@
apply assumption
apply (rule notI)
unfolding double_zero [symmetric, of a]
- apply simp
+ apply blast
done
qed
@@ -259,7 +262,8 @@
"a + a \<le> 0 \<longleftrightarrow> a \<le> 0"
proof -
have "a + a \<le> 0 \<longleftrightarrow> 0 \<le> - (a + a)" by (subst le_minus_iff, simp)
- moreover have "\<dots> \<longleftrightarrow> a \<le> 0" by simp
+ moreover have "\<dots> \<longleftrightarrow> a \<le> 0"
+ by (simp only: minus_add_distrib zero_le_double_add_iff_zero_le_single_add) simp
ultimately show ?thesis by blast
qed
@@ -267,11 +271,12 @@
"a + a < 0 \<longleftrightarrow> a < 0"
proof -
have "a + a < 0 \<longleftrightarrow> 0 < - (a + a)" by (subst less_minus_iff, simp)
- moreover have "\<dots> \<longleftrightarrow> a < 0" by simp
+ moreover have "\<dots> \<longleftrightarrow> a < 0"
+ by (simp only: minus_add_distrib zero_less_double_add_iff_zero_less_single_add) simp
ultimately show ?thesis by blast
qed
-declare neg_inf_eq_sup [simp] neg_sup_eq_inf [simp]
+declare neg_inf_eq_sup [simp] neg_sup_eq_inf [simp] diff_inf_eq_sup [simp] diff_sup_eq_inf [simp]
lemma le_minus_self_iff: "a \<le> - a \<longleftrightarrow> a \<le> 0"
proof -
@@ -326,7 +331,7 @@
then have "0 \<le> sup a (- a)" unfolding abs_lattice .
then have "sup (sup a (- a)) 0 = sup a (- a)" by (rule sup_absorb1)
then show ?thesis
- by (simp add: add_sup_inf_distribs sup_aci pprt_def nprt_def diff_minus abs_lattice)
+ by (simp add: add_sup_inf_distribs ac_simps pprt_def nprt_def abs_lattice)
qed
subclass ordered_ab_group_add_abs
@@ -355,16 +360,17 @@
show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
proof -
have g:"abs a + abs b = sup (a+b) (sup (-a-b) (sup (-a+b) (a + (-b))))" (is "_=sup ?m ?n")
- by (simp add: abs_lattice add_sup_inf_distribs sup_aci diff_minus)
+ by (simp add: abs_lattice add_sup_inf_distribs sup_aci ac_simps)
have a: "a + b <= sup ?m ?n" by simp
have b: "- a - b <= ?n" by simp
have c: "?n <= sup ?m ?n" by simp
from b c have d: "-a-b <= sup ?m ?n" by (rule order_trans)
- have e:"-a-b = -(a+b)" by (simp add: diff_minus)
+ have e:"-a-b = -(a+b)" by simp
from a d e have "abs(a+b) <= sup ?m ?n"
apply -
apply (drule abs_leI)
- apply auto
+ apply (simp_all only: algebra_simps ac_simps minus_add)
+ apply (metis add_uminus_conv_diff d sup_commute uminus_add_conv_diff)
done
with g[symmetric] show ?thesis by simp
qed
@@ -421,14 +427,12 @@
}
note b = this[OF refl[of a] refl[of b]]
have xy: "- ?x <= ?y"
- apply (simp)
- apply (rule order_trans [OF add_nonpos_nonpos add_nonneg_nonneg])
- apply (simp_all add: mult_nonneg_nonneg mult_nonpos_nonpos)
+ apply simp
+ apply (metis (full_types) add_increasing add_uminus_conv_diff lattice_ab_group_add_class.minus_le_self_iff minus_add_distrib mult_nonneg_nonneg mult_nonpos_nonpos nprt_le_zero zero_le_pprt)
done
have yx: "?y <= ?x"
- apply (simp add:diff_minus)
- apply (rule order_trans [OF add_nonpos_nonpos add_nonneg_nonneg])
- apply (simp_all add: mult_nonneg_nonpos mult_nonpos_nonneg)
+ apply simp
+ apply (metis (full_types) add_nonpos_nonpos add_uminus_conv_diff lattice_ab_group_add_class.le_minus_self_iff minus_add_distrib mult_nonneg_nonpos mult_nonpos_nonneg nprt_le_zero zero_le_pprt)
done
have i1: "a*b <= abs a * abs b" by (simp only: a b yx)
have i2: "- (abs a * abs b) <= a*b" by (simp only: a b xy)
@@ -549,7 +553,7 @@
by simp
then have "-(- nprt a1 * pprt b2 + - nprt a2 * nprt b2 + - pprt a1 * pprt b1 + - pprt a2 * nprt b1) <= a * b"
by (simp only: minus_le_iff)
- then show ?thesis by simp
+ then show ?thesis by (simp add: algebra_simps)
qed
instance int :: lattice_ring
@@ -567,3 +571,4 @@
qed
end
+
--- a/src/HOL/Library/Polynomial.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Library/Polynomial.thy Fri Nov 01 18:51:14 2013 +0100
@@ -667,7 +667,7 @@
show "- p + p = 0"
by (simp add: poly_eq_iff)
show "p - q = p + - q"
- by (simp add: poly_eq_iff diff_minus)
+ by (simp add: poly_eq_iff)
qed
end
@@ -714,15 +714,15 @@
lemma degree_diff_le_max: "degree (p - q) \<le> max (degree p) (degree q)"
using degree_add_le [where p=p and q="-q"]
- by (simp add: diff_minus)
+ by simp
lemma degree_diff_le:
"\<lbrakk>degree p \<le> n; degree q \<le> n\<rbrakk> \<Longrightarrow> degree (p - q) \<le> n"
- by (simp add: diff_minus degree_add_le)
+ using degree_add_le [of p n "- q"] by simp
lemma degree_diff_less:
"\<lbrakk>degree p < n; degree q < n\<rbrakk> \<Longrightarrow> degree (p - q) < n"
- by (simp add: diff_minus degree_add_less)
+ using degree_add_less [of p n "- q"] by simp
lemma add_monom: "monom a n + monom b n = monom (a + b) n"
by (rule poly_eqI) simp
@@ -793,7 +793,7 @@
lemma poly_diff [simp]:
fixes x :: "'a::comm_ring"
shows "poly (p - q) x = poly p x - poly q x"
- by (simp add: diff_minus)
+ using poly_add [of p "- q" x] by simp
lemma poly_setsum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)"
by (induct A rule: infinite_finite_induct) simp_all
--- a/src/HOL/Library/Product_plus.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Library/Product_plus.thy Fri Nov 01 18:51:14 2013 +0100
@@ -104,7 +104,7 @@
(cancel_comm_monoid_add, cancel_comm_monoid_add) cancel_comm_monoid_add ..
instance prod :: (group_add, group_add) group_add
- by default (simp_all add: prod_eq_iff diff_minus)
+ by default (simp_all add: prod_eq_iff)
instance prod :: (ab_group_add, ab_group_add) ab_group_add
by default (simp_all add: prod_eq_iff)
--- a/src/HOL/Library/Set_Algebras.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Library/Set_Algebras.thy Fri Nov 01 18:51:14 2013 +0100
@@ -190,12 +190,12 @@
done
lemma set_plus_imp_minus: "(a::'a::ab_group_add) : b +o C ==> (a - b) : C"
- by (auto simp add: elt_set_plus_def add_ac diff_minus)
+ by (auto simp add: elt_set_plus_def add_ac)
lemma set_minus_imp_plus: "(a::'a::ab_group_add) - b : C ==> a : b +o C"
- apply (auto simp add: elt_set_plus_def add_ac diff_minus)
+ apply (auto simp add: elt_set_plus_def add_ac)
apply (subgoal_tac "a = (a + - b) + b")
- apply (rule bexI, assumption, assumption)
+ apply (rule bexI, assumption)
apply (auto simp add: add_ac)
done
--- a/src/HOL/Limits.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Limits.thy Fri Nov 01 18:51:14 2013 +0100
@@ -179,7 +179,7 @@
apply (rule_tac x = K in exI, simp)
apply (rule exI [where x = 0], auto)
apply (erule order_less_le_trans, simp)
-apply (drule_tac x=n in spec, fold diff_minus)
+apply (drule_tac x=n in spec)
apply (drule order_trans [OF norm_triangle_ineq2])
apply simp
done
@@ -192,9 +192,11 @@
then obtain K
where *: "0 < K" and **: "\<And>n. norm (X n) \<le> K" by (auto simp add: Bseq_def)
from * have "0 < K + norm (X 0)" by (rule order_less_le_trans) simp
- moreover from ** have "\<forall>n. norm (X n + - X 0) \<le> K + norm (X 0)"
- by (auto intro: order_trans norm_triangle_ineq)
- ultimately show ?Q by blast
+ from ** have "\<forall>n. norm (X n - X 0) \<le> K + norm (X 0)"
+ by (auto intro: order_trans norm_triangle_ineq4)
+ then have "\<forall>n. norm (X n + - X 0) \<le> K + norm (X 0)"
+ by simp
+ with `0 < K + norm (X 0)` show ?Q by blast
next
assume ?Q then show ?P by (auto simp add: Bseq_iff2)
qed
@@ -205,6 +207,7 @@
apply (drule_tac x = n in spec, arith)
done
+
subsubsection{*Upper Bounds and Lubs of Bounded Sequences*}
lemma Bseq_isUb:
@@ -342,7 +345,7 @@
unfolding Zfun_def by simp
lemma Zfun_diff: "\<lbrakk>Zfun f F; Zfun g F\<rbrakk> \<Longrightarrow> Zfun (\<lambda>x. f x - g x) F"
- by (simp only: diff_minus Zfun_add Zfun_minus)
+ using Zfun_add [of f F "\<lambda>x. - g x"] by (simp add: Zfun_minus)
lemma (in bounded_linear) Zfun:
assumes g: "Zfun g F"
@@ -532,7 +535,7 @@
lemma tendsto_diff [tendsto_intros]:
fixes a b :: "'a::real_normed_vector"
shows "\<lbrakk>(f ---> a) F; (g ---> b) F\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x - g x) ---> a - b) F"
- by (simp add: diff_minus tendsto_add tendsto_minus)
+ using tendsto_add [of f a F "\<lambda>x. - g x" "- b"] by (simp add: tendsto_minus)
lemma continuous_diff [continuous_intros]:
fixes f g :: "'a::t2_space \<Rightarrow> 'b::real_normed_vector"
--- a/src/HOL/Matrix_LP/LP.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Matrix_LP/LP.thy Fri Nov 01 18:51:14 2013 +0100
@@ -72,8 +72,7 @@
proof -
have "0 <= A - A1"
proof -
- have 1: "A - A1 = A + (- A1)" by simp
- show ?thesis by (simp only: 1 add_right_mono[of A1 A "-A1", simplified, simplified assms])
+ from assms add_right_mono [of A1 A "- A1"] show ?thesis by simp
qed
then have "abs (A-A1) = A-A1" by (rule abs_of_nonneg)
with assms show "abs (A-A1) <= (A2-A1)" by simp
@@ -147,9 +146,9 @@
then have "c * x <= y * b - (y * A - c) * x" by (simp add: le_diff_eq)
then have cx: "c * x <= y * b + (c - y * A) * x" by (simp add: algebra_simps)
have s2: "c - y * A <= c2 - y * A1"
- by (simp add: diff_minus assms add_mono mult_left_mono)
+ by (simp add: assms add_mono mult_left_mono algebra_simps)
have s1: "c1 - y * A2 <= c - y * A"
- by (simp add: diff_minus assms add_mono mult_left_mono)
+ by (simp add: assms add_mono mult_left_mono algebra_simps)
have prts: "(c - y * A) * x <= ?C"
apply (simp add: Let_def)
apply (rule mult_le_prts)
--- a/src/HOL/Matrix_LP/Matrix.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Matrix_LP/Matrix.thy Fri Nov 01 18:51:14 2013 +0100
@@ -1542,8 +1542,8 @@
fix A B :: "'a matrix"
show "- A + A = 0"
by (simp add: plus_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext)
- show "A - B = A + - B"
- by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject [symmetric] ext diff_minus)
+ show "A + - B = A - B"
+ by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject [symmetric] ext)
qed
instance matrix :: (ab_group_add) ab_group_add
--- a/src/HOL/Metis_Examples/Big_O.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Metis_Examples/Big_O.thy Fri Nov 01 18:51:14 2013 +0100
@@ -493,8 +493,10 @@
lemma bigo_compose2:
"f =o g +o O(h) \<Longrightarrow> (\<lambda>x. f(k x)) =o (\<lambda>x. g(k x)) +o O(\<lambda>x. h(k x))"
-apply (simp only: set_minus_plus [symmetric] diff_minus fun_Compl_def func_plus)
-by (erule bigo_compose1)
+apply (simp only: set_minus_plus [symmetric] fun_Compl_def func_plus)
+apply (drule bigo_compose1 [of "f - g" h k])
+apply (simp add: fun_diff_def)
+done
subsection {* Setsum *}
--- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Fri Nov 01 18:51:14 2013 +0100
@@ -1099,8 +1099,8 @@
shows "(\<lambda>x. m *s x + c) o (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id"
"(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) o (\<lambda>x. m *s x + c) = id"
using m0
- apply (auto simp add: fun_eq_iff vector_add_ldistrib)
- apply (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric])
+ apply (auto simp add: fun_eq_iff vector_add_ldistrib diff_conv_add_uminus simp del: add_uminus_conv_diff)
+ apply (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1 [symmetric])
done
lemma vector_affinity_eq:
@@ -1114,7 +1114,7 @@
using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
next
assume h: "x = inverse m *s y + - (inverse m *s c)"
- show "m *s x + c = y" unfolding h diff_minus[symmetric]
+ show "m *s x + c = y" unfolding h
using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
qed
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Fri Nov 01 18:51:14 2013 +0100
@@ -858,9 +858,10 @@
assumes "affine_parallel A B"
shows "affine_parallel B A"
proof -
- from assms obtain a where "B = (\<lambda>x. a + x) ` A"
+ from assms obtain a where B: "B = (\<lambda>x. a + x) ` A"
unfolding affine_parallel_def by auto
- then show ?thesis
+ have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
+ from B show ?thesis
using translation_galois [of B a A]
unfolding affine_parallel_def by auto
qed
@@ -980,6 +981,7 @@
assumes "affine S" "a \<in> S"
shows "subspace ((\<lambda>x. (-a)+x) ` S)"
proof -
+ have [simp]: "(\<lambda>x. x - a) = plus (- a)" by (simp add: fun_eq_iff)
have "affine ((\<lambda>x. (-a)+x) ` S)"
using affine_translation assms by auto
moreover have "0 : ((\<lambda>x. (-a)+x) ` S)"
@@ -992,15 +994,12 @@
assumes "L \<equiv> {y. \<exists>x \<in> S. (-a)+x=y}"
shows "subspace L & affine_parallel S L"
proof -
- have par: "affine_parallel S L"
- unfolding affine_parallel_def using assms by auto
+ from assms have "L = plus (- a) ` S" by auto
+ then have par: "affine_parallel S L"
+ unfolding affine_parallel_def ..
then have "affine L" using assms parallel_is_affine by auto
moreover have "0 \<in> L"
- using assms
- apply auto
- using exI[of "(\<lambda>x. x:S \<and> -a+x=0)" a]
- apply auto
- done
+ using assms by auto
ultimately show ?thesis
using subspace_affine par by auto
qed
@@ -2390,7 +2389,7 @@
ultimately have h1: "affine hull ((\<lambda>x. a + x) ` S) \<subseteq> (\<lambda>x. a + x) ` (affine hull S)"
by (metis hull_minimal)
have "affine((\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S)))"
- using affine_translation affine_affine_hull by auto
+ using affine_translation affine_affine_hull by (auto simp del: uminus_add_conv_diff)
moreover have "(\<lambda>x. -a + x) ` (\<lambda>x. a + x) ` S \<subseteq> (\<lambda>x. -a + x) ` (affine hull ((\<lambda>x. a + x) ` S))"
using hull_subset[of "(\<lambda>x. a + x) ` S"] by auto
moreover have "S = (\<lambda>x. -a + x) ` (\<lambda>x. a + x) ` S"
@@ -2478,7 +2477,7 @@
using affine_dependent_translation_eq[of "(insert a S)" "-a"]
affine_dependent_imp_dependent2 assms
dependent_imp_affine_dependent[of a S]
- by auto
+ by (auto simp del: uminus_add_conv_diff)
qed
lemma affine_dependent_iff_dependent2:
@@ -2512,7 +2511,7 @@
then have "insert 0 ((\<lambda>x. -a+x) ` (s - {a})) = (\<lambda>x. -a+x) ` s"
by auto
then have "span ((\<lambda>x. -a+x) ` (s - {a}))=span ((\<lambda>x. -a+x) ` s)"
- using span_insert_0[of "op + (- a) ` (s - {a})"] by auto
+ using span_insert_0[of "op + (- a) ` (s - {a})"] by (auto simp del: uminus_add_conv_diff)
moreover have "{x - a |x. x \<in> (s - {a})} = ((\<lambda>x. -a+x) ` (s - {a}))"
by auto
moreover have "insert a (s - {a}) = insert a s"
@@ -2652,7 +2651,7 @@
moreover have h1: "card ((\<lambda>x. -a + x) ` (B-{a})) = card (B-{a})"
apply (rule card_image)
using translate_inj_on
- apply auto
+ apply (auto simp del: uminus_add_conv_diff)
done
ultimately have "card (B-{a}) > 0" by auto
then have *: "finite (B - {a})"
--- a/src/HOL/Multivariate_Analysis/Derivative.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy Fri Nov 01 18:51:14 2013 +0100
@@ -516,7 +516,7 @@
unfolding e_def
using c[THEN conjunct1]
using norm_minus_cancel[of "f' i - f'' i"]
- by (auto simp add: add.commute ab_diff_minus)
+ by auto
finally show False
using c
using d[THEN conjunct2,rule_format,of "x + c *\<^sub>R i"]
--- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Fri Nov 01 18:51:14 2013 +0100
@@ -115,7 +115,7 @@
instance vec :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add ..
instance vec :: (group_add, finite) group_add
- by default (simp_all add: vec_eq_iff diff_minus)
+ by default (simp_all add: vec_eq_iff)
instance vec :: (ab_group_add, finite) ab_group_add
by default (simp_all add: vec_eq_iff)
--- a/src/HOL/Multivariate_Analysis/Integration.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Integration.thy Fri Nov 01 18:51:14 2013 +0100
@@ -44,9 +44,7 @@
by auto
also have "\<dots> \<le> e"
apply (rule cSup_asclose)
- apply (auto simp add: S)
- apply (metis abs_minus_add_cancel b add_commute diff_minus)
- done
+ using abs_minus_add_cancel b by (auto simp add: S)
finally have "\<bar>- Sup (uminus ` S) - l\<bar> \<le> e" .
then show ?thesis
by (simp add: Inf_real_def)
@@ -380,7 +378,7 @@
using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto
then have "y\<bullet>k < a\<bullet>k"
using e[THEN conjunct1] k
- by (auto simp add: field_simps as inner_Basis inner_simps)
+ by (auto simp add: field_simps abs_less_iff as inner_Basis inner_simps)
then have "y \<notin> i"
unfolding ab mem_interval by (auto intro!: bexI[OF _ k])
then show False using yi by auto
@@ -11975,7 +11973,7 @@
and "g absolutely_integrable_on s"
shows "(\<lambda>x. f x - g x) absolutely_integrable_on s"
using absolutely_integrable_add[OF assms(1) absolutely_integrable_neg[OF assms(2)]]
- by (simp only: algebra_simps)
+ by (simp add: algebra_simps)
lemma absolutely_integrable_linear:
fixes f :: "'m::ordered_euclidean_space \<Rightarrow> 'n::ordered_euclidean_space"
--- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Fri Nov 01 18:51:14 2013 +0100
@@ -303,7 +303,7 @@
by (metis linear_iff)
lemma linear_sub: "linear f \<Longrightarrow> f (x - y) = f x - f y"
- by (simp add: diff_minus linear_add linear_neg)
+ using linear_add [of f x "- y"] by (simp add: linear_neg)
lemma linear_setsum:
assumes lin: "linear f"
@@ -384,10 +384,10 @@
using bilinear_radd [OF assms, of x 0 0 ] by (simp add: eq_add_iff field_simps)
lemma bilinear_lsub: "bilinear h \<Longrightarrow> h (x - y) z = h x z - h y z"
- by (simp add: diff_minus bilinear_ladd bilinear_lneg)
+ using bilinear_ladd [of h x "- y"] by (simp add: bilinear_lneg)
lemma bilinear_rsub: "bilinear h \<Longrightarrow> h z (x - y) = h z x - h z y"
- by (simp add: diff_minus bilinear_radd bilinear_rneg)
+ using bilinear_radd [of h _ x "- y"] by (simp add: bilinear_rneg)
lemma bilinear_setsum:
assumes bh: "bilinear h"
@@ -730,7 +730,7 @@
by (metis scaleR_minus1_left subspace_mul)
lemma subspace_sub: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x - y \<in> S"
- by (metis diff_minus subspace_add subspace_neg)
+ using subspace_add [of S x "- y"] by (simp add: subspace_neg)
lemma (in real_vector) subspace_setsum:
assumes sA: "subspace A"
@@ -1021,8 +1021,7 @@
apply safe
apply (rule_tac x=k in exI, simp)
apply (erule rev_image_eqI [OF SigmaI [OF rangeI]])
- apply simp
- apply (rule right_minus)
+ apply auto
done
then show ?thesis by simp
qed
@@ -2064,7 +2063,7 @@
using C by simp
have "orthogonal ?a y"
unfolding orthogonal_def
- unfolding inner_diff inner_setsum_left diff_eq_0_iff_eq
+ unfolding inner_diff inner_setsum_left right_minus_eq
unfolding setsum_diff1' [OF `finite C` `y \<in> C`]
apply (clarsimp simp add: inner_commute[of y a])
apply (rule setsum_0')
@@ -3026,7 +3025,7 @@
norm x * (norm y * (y \<bullet> x) - norm x * norm y * norm y) = 0)"
using x y
unfolding inner_simps
- unfolding power2_norm_eq_inner[symmetric] power2_eq_square diff_eq_0_iff_eq
+ unfolding power2_norm_eq_inner[symmetric] power2_eq_square right_minus_eq
apply (simp add: inner_commute)
apply (simp add: field_simps)
apply metis
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Fri Nov 01 18:51:14 2013 +0100
@@ -4858,8 +4858,8 @@
assumes "uniformly_continuous_on s f"
and "uniformly_continuous_on s g"
shows "uniformly_continuous_on s (\<lambda>x. f x - g x)"
- unfolding ab_diff_minus using assms
- by (intro uniformly_continuous_on_add uniformly_continuous_on_minus)
+ using assms uniformly_continuous_on_add [of s f "- g"]
+ by (simp add: fun_Compl_def uniformly_continuous_on_minus)
text{* Continuity of all kinds is preserved under composition. *}
@@ -5680,8 +5680,6 @@
apply auto
apply (rule_tac x= xa in exI)
apply auto
- apply (rule_tac x=xa in exI)
- apply auto
done
then show ?thesis
using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto
@@ -7032,7 +7030,8 @@
unfolding homeomorphic_minimal
apply (rule_tac x="\<lambda>x. a + x" in exI)
apply (rule_tac x="\<lambda>x. -a + x" in exI)
- using continuous_on_add[OF continuous_on_const continuous_on_id]
+ using continuous_on_add [OF continuous_on_const continuous_on_id, of s a]
+ continuous_on_add [OF continuous_on_const continuous_on_id, of "plus a ` s" "- a"]
apply auto
done
--- a/src/HOL/NSA/CLim.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/NSA/CLim.thy Fri Nov 01 18:51:14 2013 +0100
@@ -22,11 +22,11 @@
lemma all_shift: "(\<forall>x::'a::comm_ring_1. P x) = (\<forall>x. P (x-a))";
apply auto
apply (drule_tac x="x+a" in spec)
-apply (simp add: diff_minus add_assoc)
+apply (simp add: add_assoc)
done
lemma complex_add_minus_iff [simp]: "(x + - a = (0::complex)) = (x=a)"
-by (simp add: diff_eq_eq diff_minus [symmetric])
+by (simp add: diff_eq_eq)
lemma complex_add_eq_0_iff [iff]: "(x+y = (0::complex)) = (y = -x)"
apply auto
--- a/src/HOL/NSA/HDeriv.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/NSA/HDeriv.thy Fri Nov 01 18:51:14 2013 +0100
@@ -81,8 +81,7 @@
text{*second equivalence *}
lemma NSDERIV_NSLIM_iff2:
"(NSDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --NS> D)"
-by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff diff_minus [symmetric]
- LIM_NSLIM_iff [symmetric])
+ by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff LIM_NSLIM_iff [symmetric])
(* while we're at it! *)
@@ -120,11 +119,10 @@
hypreal_of_real (f x))\<approx> (hypreal_of_real D) * h)"
apply (auto simp add: nsderiv_def)
apply (case_tac "h = (0::hypreal) ")
-apply (auto simp add: diff_minus)
+apply auto
apply (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: diff_minus)
+apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD])
done
lemma NSDERIVD3:
@@ -135,8 +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 diff_minus)
+apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult_assoc)
done
text{*Differentiability implies continuity
@@ -174,7 +171,7 @@
apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def)
apply (auto simp add: add_divide_distrib diff_divide_distrib dest!: spec)
apply (drule_tac b = "star_of Da" and d = "star_of Db" in approx_add)
-apply (auto simp add: diff_minus add_ac)
+apply (auto simp add: add_ac algebra_simps)
done
text{*Product of functions - Proof is trivial but tedious
@@ -234,9 +231,11 @@
hence deriv: "(\<lambda>h. - ((f(x+h) - f x) / h)) -- 0 --NS> - D"
by (rule NSLIM_minus)
have "\<forall>h. - ((f (x + h) - f x) / h) = (- f (x + h) + f x) / h"
- by (simp add: minus_divide_left diff_minus)
+ by (simp add: minus_divide_left)
with deriv
- show "(\<lambda>h. (- f (x + h) + f x) / h) -- 0 --NS> - D" by simp
+ have "(\<lambda>h. (- f (x + h) + f x) / h) -- 0 --NS> - D" by simp
+ then show "(\<lambda>h. (f (x + h) - f x) / h) -- 0 --NS> D \<Longrightarrow>
+ (\<lambda>h. (f x - f (x + h)) / h) -- 0 --NS> - D" by simp
qed
text{*Subtraction*}
@@ -244,11 +243,8 @@
by (blast dest: NSDERIV_add NSDERIV_minus)
lemma NSDERIV_diff:
- "[| NSDERIV f x :> Da; NSDERIV g x :> Db |]
- ==> NSDERIV (%x. f x - g x) x :> Da-Db"
-apply (simp add: diff_minus)
-apply (blast intro: NSDERIV_add_minus)
-done
+ "NSDERIV f x :> Da \<Longrightarrow> NSDERIV g x :> Db \<Longrightarrow> NSDERIV (\<lambda>x. f x - g x) x :> Da-Db"
+ using NSDERIV_add_minus [of f x Da g Db] by simp
text{* Similarly to the above, the chain rule admits an entirely
straightforward derivation. Compare this with Harrison's
@@ -294,7 +290,7 @@
- star_of (f (g x)))
/ (( *f* g) (star_of(x) + xa) - star_of (g x))
\<approx> star_of(Da)"
-by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def diff_minus [symmetric])
+by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def)
(*--------------------------------------------------------------
from other version of differentiability
@@ -354,13 +350,23 @@
from h_Inf have "h * star_of x \<in> Infinitesimal" by (rule Infinitesimal_HFinite_mult) simp
with assms have "inverse (- (h * star_of x) + - (star_of x * star_of x)) \<approx>
inverse (- (star_of x * star_of x))"
- by (auto intro: inverse_add_Infinitesimal_approx2
+ apply - apply (rule inverse_add_Infinitesimal_approx2)
+ apply (auto
dest!: hypreal_of_real_HFinite_diff_Infinitesimal
simp add: inverse_minus_eq [symmetric] HFinite_minus_iff)
- with not_0 `h \<noteq> 0` assms have "(inverse (star_of x + h) - inverse (star_of x)) / h \<approx>
+ done
+ moreover from not_0 `h \<noteq> 0` assms
+ have "inverse (- (h * star_of x) + - (star_of x * star_of x)) =
+ (inverse (star_of x + h) - inverse (star_of x)) / h"
+ apply (simp add: division_ring_inverse_diff nonzero_inverse_mult_distrib [symmetric]
+ nonzero_inverse_minus_eq [symmetric] ac_simps ring_distribs)
+ apply (subst nonzero_inverse_minus_eq [symmetric])
+ using distrib_right [symmetric, of h "star_of x" "star_of x"] apply simp
+ apply (simp add: field_simps)
+ done
+ ultimately have "(inverse (star_of x + h) - inverse (star_of x)) / h \<approx>
- (inverse (star_of x) * inverse (star_of x))"
- by (simp add: inverse_add nonzero_inverse_mult_distrib [symmetric]
- nonzero_inverse_minus_eq [symmetric] add_ac mult_ac diff_minus ring_distribs)
+ using assms by (simp add: nonzero_inverse_mult_distrib [symmetric] nonzero_inverse_minus_eq [symmetric])
} then show ?thesis by (simp add: nsderiv_def)
qed
--- a/src/HOL/NSA/HLim.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/NSA/HLim.thy Fri Nov 01 18:51:14 2013 +0100
@@ -71,7 +71,7 @@
lemma NSLIM_diff:
"\<lbrakk>f -- x --NS> l; g -- x --NS> m\<rbrakk> \<Longrightarrow> (\<lambda>x. f x - g x) -- x --NS> (l - m)"
-by (simp only: diff_minus NSLIM_add NSLIM_minus)
+ by (simp only: NSLIM_add NSLIM_minus diff_conv_add_uminus)
lemma NSLIM_add_minus: "[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)"
by (simp only: NSLIM_add NSLIM_minus)
@@ -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: diff_minus add_assoc)
+apply (auto simp add: add_assoc)
done
lemma NSLIM_const_not_eq:
@@ -243,14 +243,14 @@
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 diff_minus [symmetric])
+ 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 approx_refl 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)"
-by (rule NSLIM_h_iff)
+ by (fact NSLIM_h_iff)
lemma isNSCont_minus: "isNSCont f a ==> isNSCont (%x. - f x) a"
by (simp add: isNSCont_def)
--- a/src/HOL/NSA/HSEQ.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/NSA/HSEQ.thy Fri Nov 01 18:51:14 2013 +0100
@@ -73,14 +73,14 @@
lemma NSLIMSEQ_minus_cancel: "(%n. -(X n)) ----NS> -a ==> X ----NS> a"
by (drule NSLIMSEQ_minus, simp)
+lemma NSLIMSEQ_diff:
+ "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n - Y n) ----NS> a - b"
+ using NSLIMSEQ_add [of X a "- Y" "- b"] by (simp add: NSLIMSEQ_minus fun_Compl_def)
+
(* FIXME: delete *)
lemma NSLIMSEQ_add_minus:
"[| X ----NS> a; Y ----NS> b |] ==> (%n. X n + -Y n) ----NS> a + -b"
-by (simp add: NSLIMSEQ_add NSLIMSEQ_minus)
-
-lemma NSLIMSEQ_diff:
- "[| X ----NS> a; Y ----NS> b |] ==> (%n. X n - Y n) ----NS> a - b"
-by (simp add: diff_minus NSLIMSEQ_add NSLIMSEQ_minus)
+ by (simp add: NSLIMSEQ_diff)
lemma NSLIMSEQ_diff_const: "f ----NS> a ==> (%n.(f n - b)) ----NS> a - b"
by (simp add: NSLIMSEQ_diff NSLIMSEQ_const)
@@ -233,11 +233,11 @@
lemma NSLIMSEQ_inverse_real_of_nat_add_minus:
"(%n. r + -inverse(real(Suc n))) ----NS> r"
-by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add_minus)
+ using LIMSEQ_inverse_real_of_nat_add_minus by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])
lemma NSLIMSEQ_inverse_real_of_nat_add_minus_mult:
"(%n. r*( 1 + -inverse(real(Suc n)))) ----NS> r"
-by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add_minus_mult)
+ using LIMSEQ_inverse_real_of_nat_add_minus_mult by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric])
subsection {* Convergence *}
--- a/src/HOL/NSA/HSeries.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/NSA/HSeries.thy Fri Nov 01 18:51:14 2013 +0100
@@ -131,7 +131,7 @@
apply (auto simp add: approx_refl)
apply (drule approx_sym [THEN approx_minus_iff [THEN iffD1]])
apply (auto dest: approx_hrabs
- simp add: sumhr_split_diff diff_minus [symmetric])
+ simp add: sumhr_split_diff)
done
(*----------------------------------------------------------------
@@ -172,7 +172,7 @@
apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym])
apply (rule_tac [2] approx_minus_iff [THEN iffD2])
apply (auto dest: approx_hrabs_zero_cancel
- simp add: sumhr_split_diff diff_minus [symmetric])
+ simp add: sumhr_split_diff)
done
--- a/src/HOL/NSA/HTranscendental.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/NSA/HTranscendental.thy Fri Nov 01 18:51:14 2013 +0100
@@ -258,7 +258,7 @@
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: diff_minus mem_infmal_iff)
+apply (auto simp add: mem_infmal_iff minus_one [symmetric] simp del: minus_one)
done
lemma STAR_exp_epsilon [simp]: "( *f* exp) epsilon @= 1"
@@ -450,7 +450,7 @@
apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
simp add: mult_assoc)
apply (rule approx_add_right_cancel [where d = "-1"])
-apply (simp add: diff_minus)
+apply (simp add: minus_one [symmetric] del: minus_one)
done
lemma STAR_tan_zero [simp]: "( *f* tan) 0 = 0"
@@ -587,7 +587,7 @@
"x \<in> Infinitesimal ==> ( *f* cos) x @= 1 - x\<^sup>2"
apply (rule STAR_cos_Infinitesimal [THEN approx_trans])
apply (auto simp add: Infinitesimal_approx_minus [symmetric]
- diff_minus add_assoc [symmetric] numeral_2_eq_2)
+ add_assoc [symmetric] numeral_2_eq_2)
done
lemma STAR_cos_Infinitesimal_approx2:
--- a/src/HOL/NSA/NSA.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/NSA/NSA.thy Fri Nov 01 18:51:14 2013 +0100
@@ -368,7 +368,7 @@
lemma Infinitesimal_diff:
"[| x \<in> Infinitesimal; y \<in> Infinitesimal |] ==> x-y \<in> Infinitesimal"
-by (simp add: diff_minus Infinitesimal_add)
+ using Infinitesimal_add [of x "- y"] by simp
lemma Infinitesimal_mult:
fixes x y :: "'a::real_normed_algebra star"
@@ -491,7 +491,9 @@
"[|(x::hypreal) \<in> HInfinite; y \<le> 0; x \<le> 0|] ==> (x + y): HInfinite"
apply (drule HInfinite_minus_iff [THEN iffD2])
apply (rule HInfinite_minus_iff [THEN iffD1])
-apply (auto intro: HInfinite_add_ge_zero)
+apply (simp only: minus_add add.commute)
+apply (rule HInfinite_add_ge_zero)
+apply simp_all
done
lemma HInfinite_add_lt_zero:
@@ -620,7 +622,7 @@
by (simp add: approx_def)
lemma approx_minus_iff2: " (x @= y) = (-y + x @= 0)"
-by (simp add: approx_def diff_minus add_commute)
+by (simp add: approx_def add_commute)
lemma approx_refl [iff]: "x @= x"
by (simp add: approx_def Infinitesimal_def)
@@ -637,7 +639,7 @@
lemma approx_trans: "[| x @= y; y @= z |] ==> x @= z"
apply (simp add: approx_def)
apply (drule (1) Infinitesimal_add)
-apply (simp add: diff_minus)
+apply simp
done
lemma approx_trans2: "[| r @= x; s @= x |] ==> r @= s"
@@ -687,7 +689,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 diff_minus)
+apply (simp add: add_commute)
done
lemma approx_minus2: "-a @= -b ==> a @= b"
@@ -700,7 +702,7 @@
by (blast intro!: approx_add approx_minus)
lemma approx_diff: "[| a @= b; c @= d |] ==> a - c @= b - d"
-by (simp only: diff_minus approx_add approx_minus)
+ using approx_add [of a b "- c" "- d"] by simp
lemma approx_mult1:
fixes a b c :: "'a::real_normed_algebra star"
@@ -1213,7 +1215,9 @@
r \<in> Reals; 0 < r |]
==> -(x + -t) \<le> r"
apply (subgoal_tac "(t + -r \<le> x)")
-apply (auto intro: lemma_st_part_le2)
+apply simp
+apply (rule lemma_st_part_le2)
+apply auto
done
lemma lemma_SReal_ub:
@@ -1238,7 +1242,7 @@
==> x + -t \<noteq> r"
apply auto
apply (frule isLubD1a [THEN Reals_minus])
-apply (drule Reals_add_cancel, assumption)
+using Reals_add_cancel [of x "- t"] apply simp
apply (drule_tac x = x in lemma_SReal_lub)
apply (drule hypreal_isLub_unique, assumption, auto)
done
@@ -1250,8 +1254,7 @@
==> -(x + -t) \<noteq> r"
apply (auto)
apply (frule isLubD1a)
-apply (drule Reals_add_cancel, assumption)
-apply (drule_tac a = "-x" in Reals_minus, simp)
+using Reals_add_cancel [of "- x" t] apply simp
apply (drule_tac x = x in lemma_SReal_lub)
apply (drule hypreal_isLub_unique, assumption, auto)
done
--- a/src/HOL/NSA/NSCA.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/NSA/NSCA.thy Fri Nov 01 18:51:14 2013 +0100
@@ -165,7 +165,7 @@
lemma approx_hcmod_approx_zero: "(x @= y) = (hcmod (y - x) @= 0)"
apply (subst hnorm_minus_commute)
-apply (simp add: approx_def Infinitesimal_hcmod_iff diff_minus)
+apply (simp add: approx_def Infinitesimal_hcmod_iff)
done
lemma approx_approx_zero_iff: "(x @= 0) = (hcmod x @= 0)"
@@ -178,14 +178,14 @@
"u @= 0 ==> hcmod(x + u) - hcmod x \<in> Infinitesimal"
apply (drule approx_approx_zero_iff [THEN iffD1])
apply (rule_tac e = "hcmod u" and e' = "- hcmod u" in Infinitesimal_interval2)
-apply (auto simp add: mem_infmal_iff [symmetric] diff_minus)
+apply (auto simp add: mem_infmal_iff [symmetric])
apply (rule_tac c1 = "hcmod x" in add_le_cancel_left [THEN iffD1])
-apply (auto simp add: diff_minus [symmetric])
+apply auto
done
lemma approx_hcmod_add_hcmod: "u @= 0 ==> hcmod(x + u) @= hcmod x"
apply (rule approx_minus_iff [THEN iffD2])
-apply (auto intro: Infinitesimal_hcmod_add_diff simp add: mem_infmal_iff [symmetric] diff_minus [symmetric])
+apply (auto intro: Infinitesimal_hcmod_add_diff simp add: mem_infmal_iff [symmetric])
done
--- a/src/HOL/NSA/StarDef.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/NSA/StarDef.thy Fri Nov 01 18:51:14 2013 +0100
@@ -803,7 +803,7 @@
instance star :: (ab_group_add) ab_group_add
apply (intro_classes)
apply (transfer, rule left_minus)
-apply (transfer, rule diff_minus)
+apply (transfer, rule diff_conv_add_uminus)
done
instance star :: (ordered_ab_semigroup_add) ordered_ab_semigroup_add
--- a/src/HOL/Num.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Num.thy Fri Nov 01 18:51:14 2013 +0100
@@ -407,7 +407,7 @@
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 minus_add_cancel add_minus_cancel)
+ apply (simp add: add_assoc)
apply (simp add: add_assoc, simp add: add_assoc [symmetric])
done
@@ -418,7 +418,7 @@
lemmas is_num_normalize =
add_assoc is_num_add_commute is_num_add_left_commute
is_num.intros is_num_numeral
- diff_minus minus_add add_minus_cancel minus_add_cancel
+ minus_add
definition dbl :: "'a \<Rightarrow> 'a" where "dbl x = x + x"
definition dbl_inc :: "'a \<Rightarrow> 'a" where "dbl_inc x = x + x + 1"
@@ -435,24 +435,21 @@
"dbl 0 = 0"
"dbl 1 = 2"
"dbl (numeral k) = numeral (Bit0 k)"
- unfolding dbl_def neg_numeral_def numeral.simps
- by (simp_all add: minus_add)
+ by (simp_all add: dbl_def neg_numeral_def numeral.simps minus_add)
lemma dbl_inc_simps [simp]:
"dbl_inc (neg_numeral k) = neg_numeral (BitM k)"
"dbl_inc 0 = 1"
"dbl_inc 1 = 3"
"dbl_inc (numeral k) = numeral (Bit1 k)"
- unfolding dbl_inc_def neg_numeral_def numeral.simps numeral_BitM
- by (simp_all add: is_num_normalize)
+ by (simp_all add: dbl_inc_def neg_numeral_def numeral.simps numeral_BitM is_num_normalize algebra_simps del: add_uminus_conv_diff)
lemma dbl_dec_simps [simp]:
"dbl_dec (neg_numeral k) = neg_numeral (Bit1 k)"
"dbl_dec 0 = -1"
"dbl_dec 1 = 1"
"dbl_dec (numeral k) = numeral (BitM k)"
- unfolding dbl_dec_def neg_numeral_def numeral.simps numeral_BitM
- by (simp_all add: is_num_normalize)
+ by (simp_all add: dbl_dec_def neg_numeral_def numeral.simps numeral_BitM is_num_normalize)
lemma sub_num_simps [simp]:
"sub One One = 0"
@@ -464,38 +461,35 @@
"sub (Bit0 k) (Bit1 l) = dbl_dec (sub k l)"
"sub (Bit1 k) (Bit0 l) = dbl_inc (sub k l)"
"sub (Bit1 k) (Bit1 l) = dbl (sub k l)"
- unfolding dbl_def dbl_dec_def dbl_inc_def sub_def
- unfolding neg_numeral_def numeral.simps numeral_BitM
- by (simp_all add: is_num_normalize)
+ by (simp_all add: dbl_def dbl_dec_def dbl_inc_def sub_def neg_numeral_def numeral.simps
+ numeral_BitM is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus)
lemma add_neg_numeral_simps:
"numeral m + neg_numeral n = sub m n"
"neg_numeral m + numeral n = sub n m"
"neg_numeral m + neg_numeral n = neg_numeral (m + n)"
- unfolding sub_def diff_minus neg_numeral_def numeral_add numeral.simps
- by (simp_all add: is_num_normalize)
+ by (simp_all add: sub_def neg_numeral_def numeral_add numeral.simps is_num_normalize
+ del: add_uminus_conv_diff add: diff_conv_add_uminus)
lemma add_neg_numeral_special:
"1 + neg_numeral m = sub One m"
"neg_numeral m + 1 = sub One m"
- unfolding sub_def diff_minus neg_numeral_def numeral_add numeral.simps
- by (simp_all add: is_num_normalize)
+ by (simp_all add: sub_def neg_numeral_def numeral_add numeral.simps is_num_normalize)
lemma diff_numeral_simps:
"numeral m - numeral n = sub m n"
"numeral m - neg_numeral n = numeral (m + n)"
"neg_numeral m - numeral n = neg_numeral (m + n)"
"neg_numeral m - neg_numeral n = sub n m"
- unfolding neg_numeral_def sub_def diff_minus numeral_add numeral.simps
- by (simp_all add: is_num_normalize)
+ by (simp_all add: neg_numeral_def sub_def numeral_add numeral.simps is_num_normalize
+ del: add_uminus_conv_diff add: diff_conv_add_uminus)
lemma diff_numeral_special:
"1 - numeral n = sub One n"
"1 - neg_numeral n = numeral (One + n)"
"numeral m - 1 = sub m One"
"neg_numeral m - 1 = neg_numeral (m + One)"
- unfolding neg_numeral_def sub_def diff_minus numeral_add numeral.simps
- by (simp_all add: is_num_normalize)
+ by (simp_all add: neg_numeral_def sub_def numeral_add numeral.simps add: is_num_normalize)
lemma minus_one: "- 1 = -1"
unfolding neg_numeral_def numeral.simps ..
--- a/src/HOL/Number_Theory/Cong.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Number_Theory/Cong.thy Fri Nov 01 18:51:14 2013 +0100
@@ -543,7 +543,8 @@
apply (subgoal_tac "a * b = (-a * -b)")
apply (erule ssubst)
apply (subst zmod_zmult2_eq)
- apply (auto simp add: mod_add_left_eq)
+ apply (auto simp add: mod_add_left_eq mod_minus_right div_minus_right)
+ apply (metis mod_diff_left_eq mod_diff_right_eq mod_mult_self1_is_0 semiring_numeral_div_class.diff_zero)+
done
lemma cong_to_1_nat: "([(a::nat) = 1] (mod n)) \<Longrightarrow> (n dvd (a - 1))"
--- a/src/HOL/Probability/Borel_Space.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Probability/Borel_Space.thy Fri Nov 01 18:51:14 2013 +0100
@@ -677,7 +677,7 @@
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
- unfolding diff_minus using assms by simp
+ using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def)
lemma borel_measurable_times[measurable (raw)]:
fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_algebra}"
@@ -719,7 +719,8 @@
proof cases
assume "b \<noteq> 0"
with `open S` have "open ((\<lambda>x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S")
- by (auto intro!: open_affinity simp: scaleR_add_right)
+ using open_affinity [of S "inverse b" "- a /\<^sub>R b"]
+ by (auto simp: algebra_simps)
hence "?S \<in> sets borel" by auto
moreover
from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"
--- a/src/HOL/Probability/Lebesgue_Integration.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Probability/Lebesgue_Integration.thy Fri Nov 01 18:51:14 2013 +0100
@@ -1528,7 +1528,7 @@
using mono by auto
ultimately show ?thesis using fg
by (auto intro!: add_mono positive_integral_mono_AE real_of_ereal_positive_mono
- simp: positive_integral_positive lebesgue_integral_def diff_minus)
+ simp: positive_integral_positive lebesgue_integral_def algebra_simps)
qed
lemma integral_mono:
@@ -1732,7 +1732,7 @@
shows "integrable M (\<lambda>t. f t - g t)"
and "(\<integral> t. f t - g t \<partial>M) = integral\<^sup>L M f - integral\<^sup>L M g"
using integral_add[OF f integral_minus(1)[OF g]]
- unfolding diff_minus integral_minus(2)[OF g]
+ unfolding integral_minus(2)[OF g]
by auto
lemma integral_indicator[simp, intro]:
--- a/src/HOL/Rat.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Rat.thy Fri Nov 01 18:51:14 2013 +0100
@@ -468,7 +468,7 @@
unfolding less_eq_rat_def less_rat_def
by (auto, drule (1) positive_add, simp add: positive_zero)
show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
- unfolding less_eq_rat_def less_rat_def by (auto simp: diff_minus)
+ unfolding less_eq_rat_def less_rat_def by auto
show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
by (rule sgn_rat_def)
show "a \<le> b \<or> b \<le> a"
@@ -665,7 +665,7 @@
by transfer simp
lemma of_rat_diff: "of_rat (a - b) = of_rat a - of_rat b"
-by (simp only: diff_minus of_rat_add of_rat_minus)
+ using of_rat_add [of a "- b"] by (simp add: of_rat_minus)
lemma of_rat_mult: "of_rat (a * b) = of_rat a * of_rat b"
apply transfer
--- a/src/HOL/Real.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Real.thy Fri Nov 01 18:51:14 2013 +0100
@@ -98,7 +98,7 @@
lemma vanishes_diff:
assumes X: "vanishes X" and Y: "vanishes Y"
shows "vanishes (\<lambda>n. X n - Y n)"
-unfolding diff_minus by (intro vanishes_add vanishes_minus X Y)
+ unfolding diff_conv_add_uminus by (intro vanishes_add vanishes_minus X Y)
lemma vanishes_mult_bounded:
assumes X: "\<exists>a>0. \<forall>n. \<bar>X n\<bar> < a"
@@ -170,7 +170,7 @@
lemma cauchy_diff [simp]:
assumes X: "cauchy X" and Y: "cauchy Y"
shows "cauchy (\<lambda>n. X n - Y n)"
-using assms unfolding diff_minus by simp
+ using assms unfolding diff_conv_add_uminus by (simp del: add_uminus_conv_diff)
lemma cauchy_imp_bounded:
assumes "cauchy X" shows "\<exists>b>0. \<forall>n. \<bar>X n\<bar> < b"
@@ -456,7 +456,7 @@
lemma diff_Real:
assumes X: "cauchy X" and Y: "cauchy Y"
shows "Real X - Real Y = Real (\<lambda>n. X n - Y n)"
- unfolding minus_real_def diff_minus
+ unfolding minus_real_def
by (simp add: minus_Real add_Real X Y)
lemma mult_Real:
@@ -607,7 +607,7 @@
unfolding less_eq_real_def less_real_def
by (auto, drule (1) positive_add, simp add: positive_zero)
show "a \<le> b \<Longrightarrow> c + a \<le> c + b"
- unfolding less_eq_real_def less_real_def by (auto simp: diff_minus) (* by auto *)
+ unfolding less_eq_real_def less_real_def by auto
(* FIXME: Procedure int_combine_numerals: c + b - (c + a) \<equiv> b + - a *)
(* Should produce c + b - (c + a) \<equiv> b - a *)
show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)"
--- a/src/HOL/Real_Vector_Spaces.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Real_Vector_Spaces.thy Fri Nov 01 18:51:14 2013 +0100
@@ -31,7 +31,7 @@
qed
lemma diff: "f (x - y) = f x - f y"
-by (simp add: add minus diff_minus)
+ using add [of x "- y"] by (simp add: minus)
lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
apply (cases "finite A")
@@ -553,8 +553,7 @@
proof -
have "norm (a + - b) \<le> norm a + norm (- b)"
by (rule norm_triangle_ineq)
- thus ?thesis
- by (simp only: diff_minus norm_minus_cancel)
+ then show ?thesis by simp
qed
lemma norm_diff_ineq:
@@ -571,7 +570,7 @@
shows "norm ((a + b) - (c + d)) \<le> norm (a - c) + norm (b - d)"
proof -
have "norm ((a + b) - (c + d)) = norm ((a - c) + (b - d))"
- by (simp add: diff_minus add_ac)
+ by (simp add: algebra_simps)
also have "\<dots> \<le> norm (a - c) + norm (b - d)"
by (rule norm_triangle_ineq)
finally show ?thesis .
--- a/src/HOL/Rings.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Rings.thy Fri Nov 01 18:51:14 2013 +0100
@@ -255,11 +255,13 @@
lemma minus_mult_commute: "- a * b = a * - b"
by simp
-lemma right_diff_distrib[algebra_simps, field_simps]: "a * (b - c) = a * b - a * c"
-by (simp add: distrib_left diff_minus)
+lemma right_diff_distrib [algebra_simps, field_simps]:
+ "a * (b - c) = a * b - a * c"
+ using distrib_left [of a b "-c "] by simp
-lemma left_diff_distrib[algebra_simps, field_simps]: "(a - b) * c = a * c - b * c"
-by (simp add: distrib_right diff_minus)
+lemma left_diff_distrib [algebra_simps, field_simps]:
+ "(a - b) * c = a * c - b * c"
+ using distrib_right [of a "- b" c] by simp
lemmas ring_distribs =
distrib_left distrib_right left_diff_distrib right_diff_distrib
@@ -331,8 +333,9 @@
then show "- x dvd y" ..
qed
-lemma dvd_diff[simp]: "x dvd y \<Longrightarrow> x dvd z \<Longrightarrow> x dvd (y - z)"
-by (simp only: diff_minus dvd_add dvd_minus_iff)
+lemma dvd_diff [simp]:
+ "x dvd y \<Longrightarrow> x dvd z \<Longrightarrow> x dvd (y - z)"
+ using dvd_add [of x y "- z"] by simp
end
@@ -755,9 +758,7 @@
proof
fix a b
show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
- by (auto simp add: abs_if not_less)
- (auto simp del: minus_add_distrib simp add: minus_add_distrib [symmetric],
- auto intro!: less_imp_le add_neg_neg)
+ by (auto simp add: abs_if not_le not_less algebra_simps simp del: add.commute dest: add_neg_neg add_nonneg_nonneg)
qed (auto simp add: abs_if)
lemma zero_le_square [simp]: "0 \<le> a * a"
--- a/src/HOL/Semiring_Normalization.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Semiring_Normalization.thy Fri Nov 01 18:51:14 2013 +0100
@@ -137,7 +137,7 @@
lemma normalizing_ring_rules:
"- x = (- 1) * x"
"x - y = x + (- y)"
- by (simp_all add: diff_minus)
+ by simp_all
lemmas normalizing_comm_ring_1_axioms =
comm_ring_1_axioms [normalizer
--- a/src/HOL/Series.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Series.thy Fri Nov 01 18:51:14 2013 +0100
@@ -34,7 +34,7 @@
lemma sumr_diff_mult_const:
"setsum f {0..<n} - (real n*r) = setsum (%i. f i - r) {0..<n::nat}"
-by (simp add: diff_minus setsum_addf real_of_nat_def)
+ by (simp add: setsum_subtractf real_of_nat_def)
lemma real_setsum_nat_ivl_bounded:
"(!!p. p < n \<Longrightarrow> f(p) \<le> K)
--- a/src/HOL/Tools/group_cancel.ML Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Tools/group_cancel.ML Fri Nov 01 18:51:14 2013 +0100
@@ -25,7 +25,7 @@
val sub1 = @{lemma "(A::'a::ab_group_add) == k + a ==> A - b == k + (a - b)"
by (simp only: add_diff_eq)}
val sub2 = @{lemma "(B::'a::ab_group_add) == k + b ==> a - B == - k + (a - b)"
- by (simp only: diff_minus minus_add add_ac)}
+ by (simp only: minus_add diff_conv_add_uminus add_ac)}
val neg1 = @{lemma "(A::'a::ab_group_add) == k + a ==> - A == - k + - a"
by (simp only: minus_add_distrib)}
val rule0 = @{lemma "(a::'a::comm_monoid_add) == a + 0"
--- a/src/HOL/Tools/numeral_simprocs.ML Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Tools/numeral_simprocs.ML Fri Nov 01 18:51:14 2013 +0100
@@ -220,7 +220,7 @@
val minus_simps = [@{thm minus_zero}, @{thm minus_one}, @{thm minus_numeral}, @{thm minus_neg_numeral}];
(*To let us treat subtraction as addition*)
-val diff_simps = [@{thm diff_minus}, @{thm minus_add_distrib}, @{thm minus_minus}];
+val diff_simps = [@{thm diff_conv_add_uminus}, @{thm minus_add_distrib}, @{thm minus_minus}];
(*To let us treat division as multiplication*)
val divide_simps = [@{thm divide_inverse}, @{thm inverse_mult_distrib}, @{thm inverse_inverse_eq}];
@@ -719,7 +719,7 @@
@{thm "times_divide_eq_left"}, @{thm "times_divide_eq_right"},
@{thm "times_divide_times_eq"},
@{thm "divide_divide_eq_right"},
- @{thm "diff_minus"}, @{thm "minus_divide_left"},
+ @{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}))))
--- a/src/HOL/Transcendental.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Transcendental.thy Fri Nov 01 18:51:14 2013 +0100
@@ -453,7 +453,7 @@
apply simp
apply (simp only: lemma_termdiff1 setsum_right_distrib)
apply (rule setsum_cong [OF refl])
- apply (simp add: diff_minus [symmetric] less_iff_Suc_add)
+ apply (simp add: less_iff_Suc_add)
apply (clarify)
apply (simp add: setsum_right_distrib lemma_realpow_diff_sumr2 mult_ac
del: setsum_op_ivl_Suc power_Suc)
@@ -1129,8 +1129,7 @@
by (rule inverse_unique [symmetric], simp add: mult_exp_exp)
lemma exp_diff: "exp (x - y) = exp x / exp y"
- unfolding diff_minus divide_inverse
- by (simp add: exp_add exp_minus)
+ using exp_add [of x "- y"] by (simp add: exp_minus divide_inverse)
subsubsection {* Properties of the Exponential Function on Reals *}
@@ -2410,13 +2409,13 @@
using sin_cos_minus_lemma [where x=x] by simp
lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y"
- by (simp add: diff_minus sin_add)
+ 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)
lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y"
- by (simp add: diff_minus cos_add)
+ 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)
@@ -2526,8 +2525,9 @@
by (simp add: inverse_eq_divide less_divide_eq)
}
note *** = this
+ have [simp]: "\<And>x y::real. 0 < x - y \<longleftrightarrow> y < x" by arith
from ** show ?thesis by (rule sumr_pos_lt_pair)
- (simp add: divide_inverse real_0_less_add_iff mult_assoc [symmetric] ***)
+ (simp add: divide_inverse mult_assoc [symmetric] ***)
qed
ultimately have "0 < (\<Sum>n. - (-1 ^ n * 2 ^ (2 * n) / real (fact (2 * n))))"
by (rule order_less_trans)
@@ -2810,7 +2810,7 @@
apply (cut_tac x="-x" in cos_zero_lemma, simp, simp)
apply (force simp add: minus_equation_iff [of x])
apply (auto simp only: odd_Suc_mult_two_ex real_of_nat_Suc distrib_right)
-apply (auto simp add: cos_add)
+apply (auto simp add: cos_diff cos_add)
done
(* ditto: but to a lesser extent *)
@@ -3833,7 +3833,7 @@
by (rule DERIV_arctan_suminf[OF `0 < r` `r < 1` `\<bar>x\<bar> < r`])
from DERIV_add_minus[OF this DERIV_arctan]
show "DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0"
- unfolding diff_minus by auto
+ by auto
qed
hence DERIV_in_rball: "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0"
using `-r < a` `b < r` by auto
@@ -3922,9 +3922,10 @@
}
hence "\<forall> x \<in> { 0 <..< 1 }. 0 \<le> ?a x n - ?diff x n" by auto
moreover have "\<And>x. isCont (\<lambda> x. ?a x n - ?diff x n) x"
- unfolding diff_minus divide_inverse
+ unfolding diff_conv_add_uminus divide_inverse
by (auto intro!: isCont_add isCont_rabs isCont_ident isCont_minus isCont_arctan
- isCont_inverse isCont_mult isCont_power isCont_const isCont_setsum)
+ isCont_inverse isCont_mult isCont_power isCont_const isCont_setsum
+ simp del: add_uminus_conv_diff)
ultimately have "0 \<le> ?a 1 n - ?diff 1 n"
by (rule LIM_less_bound)
hence "?diff 1 n \<le> ?a 1 n" by auto
--- a/src/HOL/Word/Bit_Representation.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Word/Bit_Representation.thy Fri Nov 01 18:51:14 2013 +0100
@@ -636,12 +636,12 @@
unfolding no_sbintr_alt2 by (drule sb_inc_lem') simp
lemma sb_dec_lem:
- "(0::int) <= - (2^k) + a ==> (a + 2^k) mod (2 * 2 ^ k) <= - (2 ^ k) + a"
- by (rule int_mod_le' [where n = "2 ^ (Suc k)" and b = "a + 2 ^ k", simplified])
+ "(0::int) \<le> - (2 ^ k) + a \<Longrightarrow> (a + 2 ^ k) mod (2 * 2 ^ k) \<le> - (2 ^ k) + a"
+ using int_mod_le'[where n = "2 ^ (Suc k)" and b = "a + 2 ^ k"] by simp
lemma sb_dec_lem':
- "(2::int) ^ k <= a ==> (a + 2 ^ k) mod (2 * 2 ^ k) <= - (2 ^ k) + a"
- by (rule iffD1 [OF diff_le_eq', THEN sb_dec_lem, simplified])
+ "(2::int) ^ k \<le> a \<Longrightarrow> (a + 2 ^ k) mod (2 * 2 ^ k) \<le> - (2 ^ k) + a"
+ by (rule sb_dec_lem) simp
lemma sbintrunc_dec:
"x >= (2 ^ n) ==> x - 2 ^ (Suc n) >= sbintrunc n x"
--- a/src/HOL/Word/WordBitwise.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/Word/WordBitwise.thy Fri Nov 01 18:51:14 2013 +0100
@@ -65,7 +65,7 @@
lemma bl_word_sub:
"to_bl (x - y) = to_bl (x + (- y))"
- by (simp add: diff_def)
+ by simp
lemma rbl_word_1:
"rev (to_bl (1 :: ('a :: len0) word))
--- a/src/HOL/ex/Dedekind_Real.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/ex/Dedekind_Real.thy Fri Nov 01 18:51:14 2013 +0100
@@ -1520,14 +1520,14 @@
have "z + x - (z + y) = (z + -z) + (x - y)"
by (simp add: algebra_simps)
with le show ?thesis
- by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"] diff_minus)
+ by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"])
qed
lemma real_sum_gt_zero_less: "(0 < S + (-W::real)) ==> (W < S)"
-by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
+by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S])
lemma real_less_sum_gt_zero: "(W < S) ==> (0 < S + (-W::real))"
-by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S] diff_minus)
+by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of S])
lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y"
apply (cases x, cases y)
@@ -1543,7 +1543,7 @@
apply (rule real_sum_gt_zero_less)
apply (drule real_less_sum_gt_zero [of x y])
apply (drule real_mult_order, assumption)
-apply (simp add: distrib_left)
+apply (simp add: algebra_simps)
done
instantiation real :: distrib_lattice
--- a/src/HOL/ex/Gauge_Integration.thy Thu Oct 31 16:54:22 2013 +0100
+++ b/src/HOL/ex/Gauge_Integration.thy Fri Nov 01 18:51:14 2013 +0100
@@ -511,9 +511,9 @@
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 (simp add: distrib_right diff_minus)
+ apply (simp add: left_diff_distrib)
apply (simp add: mult_assoc divide_inverse)
- apply (simp add: distrib_right)
+ apply (simp add: ring_distribs)
done
moreover from False `\<bar>z - x\<bar> < s` have "\<bar>(f z - f x) / (z - x) - f' x\<bar> < e / 2"
by (rule P)