Skip to content

Commit db5a937

Browse files
committed
chore: Remove unnecessary "rw"s (#10704)
Remove unnecessary "rw"s.
1 parent c76aa9d commit db5a937

File tree

41 files changed

+46
-53
lines changed

Some content is hidden

Large Commits have some content hidden by default. Use the searchbox below for content that may be hidden.

41 files changed

+46
-53
lines changed

Mathlib/Algebra/Algebra/Operations.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -596,7 +596,7 @@ def span.ringHom : SetSemiring A →+* Submodule R A where
596596
map_add' := span_union
597597
map_mul' s t := by
598598
dsimp only -- porting note: new, needed due to new-style structures
599-
rw [SetSemiring.down_mul, span_mul_span, ← image_mul_prod]
599+
rw [SetSemiring.down_mul, span_mul_span]
600600
#align submodule.span.ring_hom Submodule.span.ringHom
601601

602602
section

Mathlib/Algebra/BigOperators/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -2174,7 +2174,7 @@ theorem prod_eq_zero_iff : ∏ x in s, f x = 0 ↔ ∃ a ∈ s, f a = 0 := by
21742174
classical
21752175
induction' s using Finset.induction_on with a s ha ih
21762176
· exact ⟨Not.elim one_ne_zero, fun ⟨_, H, _⟩ => by simp at H⟩
2177-
· rw [prod_insert ha, mul_eq_zero, exists_mem_insert, ih, ← bex_def]
2177+
· rw [prod_insert ha, mul_eq_zero, exists_mem_insert, ih]
21782178
#align finset.prod_eq_zero_iff Finset.prod_eq_zero_iff
21792179

21802180
theorem prod_ne_zero_iff : ∏ x in s, f x ≠ 0 ↔ ∀ a ∈ s, f a ≠ 0 := by

Mathlib/Algebra/DirectSum/Algebra.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -68,7 +68,7 @@ instance _root_.GradedMonoid.isScalarTower_right :
6868
IsScalarTower R (GradedMonoid A) (GradedMonoid A) where
6969
smul_assoc s x y := by
7070
dsimp
71-
rw [GAlgebra.smul_def, GAlgebra.smul_def, ← mul_assoc, GAlgebra.commutes, mul_assoc]
71+
rw [GAlgebra.smul_def, GAlgebra.smul_def, ← mul_assoc]
7272

7373
instance : Algebra R (⨁ i, A i) where
7474
toFun := (DirectSum.of A 0).comp GAlgebra.toFun

Mathlib/Algebra/GCDMonoid/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1145,7 +1145,7 @@ noncomputable def gcdMonoidOfLCM [DecidableEq α] (lcm : α → α → α)
11451145
dsimp only
11461146
split_ifs with h h_1
11471147
· rw [h, eq_zero_of_zero_dvd (dvd_lcm_left _ _), mul_zero, zero_mul]
1148-
· rw [h_1, eq_zero_of_zero_dvd (dvd_lcm_right _ _), mul_zero]
1148+
· rw [h_1, eq_zero_of_zero_dvd (dvd_lcm_right _ _)]
11491149
rw [mul_comm, ← Classical.choose_spec (exists_gcd a b)]
11501150
lcm_zero_left := fun a => eq_zero_of_zero_dvd (dvd_lcm_left _ _)
11511151
lcm_zero_right := fun a => eq_zero_of_zero_dvd (dvd_lcm_right _ _)

Mathlib/Algebra/Quandle.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -163,7 +163,7 @@ lemma act_act_self_eq (x y : S) : (x ◃ y) ◃ x = x ◃ y := by
163163
rw [h, ← Shelf.self_distrib, act_one]
164164
#align unital_shelf.act_act_self_eq UnitalShelf.act_act_self_eq
165165

166-
lemma act_idem (x : S) : (x ◃ x) = x := by rw [← act_one x, ← Shelf.self_distrib, act_one, act_one]
166+
lemma act_idem (x : S) : (x ◃ x) = x := by rw [← act_one x, ← Shelf.self_distrib, act_one]
167167
#align unital_shelf.act_idem UnitalShelf.act_idem
168168

169169
lemma act_self_act_eq (x y : S) : x ◃ (x ◃ y) = x ◃ y := by

Mathlib/Algebra/Regular/Basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -280,11 +280,11 @@ theorem not_isRegular_zero [Nontrivial R] : ¬IsRegular (0 : R) := fun h => IsRe
280280

281281
@[simp] lemma IsLeftRegular.mul_left_eq_zero_iff (hb : IsLeftRegular b) : b * a = 0 ↔ a = 0 := by
282282
nth_rw 1 [← mul_zero b]
283-
exact ⟨fun h ↦ hb h, fun ha ↦ by rw [ha, mul_zero]⟩
283+
exact ⟨fun h ↦ hb h, fun ha ↦ by rw [ha]⟩
284284

285285
@[simp] lemma IsRightRegular.mul_right_eq_zero_iff (hb : IsRightRegular b) : a * b = 0 ↔ a = 0 := by
286286
nth_rw 1 [← zero_mul b]
287-
exact ⟨fun h ↦ hb h, fun ha ↦ by rw [ha, zero_mul]⟩
287+
exact ⟨fun h ↦ hb h, fun ha ↦ by rw [ha]⟩
288288

289289
end MulZeroClass
290290

Mathlib/AlgebraicGeometry/OpenImmersion.lean

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -594,8 +594,7 @@ theorem range_pullback_to_base_of_left :
594594
Set.range f.1.base ∩ Set.range g.1.base := by
595595
rw [pullback.condition, Scheme.comp_val_base, coe_comp, Set.range_comp,
596596
range_pullback_snd_of_left, Opens.carrier_eq_coe,
597-
Opens.map_obj, Opens.coe_mk, Set.image_preimage_eq_inter_range,
598-
Set.inter_comm]
597+
Opens.map_obj, Opens.coe_mk, Set.image_preimage_eq_inter_range]
599598
#align algebraic_geometry.IsOpenImmersion.range_pullback_to_base_of_left AlgebraicGeometry.IsOpenImmersion.range_pullback_to_base_of_left
600599

601600
theorem range_pullback_to_base_of_right :

Mathlib/Analysis/BoxIntegral/Partition/Split.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -66,7 +66,7 @@ theorem coe_splitLower : (splitLower I i x : Set (ι → ℝ)) = ↑I ∩ { y |
6666
ext y
6767
simp only [mem_univ_pi, mem_Ioc, mem_inter_iff, mem_coe, mem_setOf_eq, forall_and, ← Pi.le_def,
6868
le_update_iff, le_min_iff, and_assoc, and_forall_ne (p := fun j => y j ≤ upper I j) i, mem_def]
69-
rw [and_comm (a := y i ≤ x), Pi.le_def]
69+
rw [and_comm (a := y i ≤ x)]
7070
#align box_integral.box.coe_split_lower BoxIntegral.Box.coe_splitLower
7171

7272
theorem splitLower_le : I.splitLower i x ≤ I :=

Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -268,8 +268,7 @@ theorem tendsto_integral_exp_smul_cocompact_of_inner_product (μ : Measure V) [
268268
have : (fun w : V →L[ℝ] ℝ => ∫ v, e[-w v] • f v) = (fun w : V => ∫ v, e[-⟪v, w⟫] • f v) ∘ A := by
269269
ext1 w
270270
congr 1 with v : 1
271-
rw [← inner_conj_symm, IsROrC.conj_to_real, InnerProductSpace.toDual_symm_apply,
272-
Real.fourierChar_apply]
271+
rw [← inner_conj_symm, IsROrC.conj_to_real, InnerProductSpace.toDual_symm_apply]
273272
rw [this]
274273
exact
275274
(tendsto_integral_exp_inner_smul_cocompact f).comp

Mathlib/Analysis/NormedSpace/ContinuousAffineMap.lean

Lines changed: 1 addition & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -260,9 +260,7 @@ def toConstProdContinuousLinearMap : (V →A[𝕜] W) ≃ₗᵢ[𝕜] W × (V
260260
left_inv f := by
261261
ext
262262
rw [f.decomp]
263-
-- Porting note: previously `simp` closed the goal, but now we need to rewrite:
264-
simp only [coe_add, ContinuousLinearMap.coe_toContinuousAffineMap, Pi.add_apply]
265-
rw [ContinuousAffineMap.coe_const, Function.const_apply]
263+
simp only [coe_add, ContinuousLinearMap.coe_toContinuousAffineMap, Pi.add_apply, coe_const]
266264
right_inv := by rintro ⟨v, f⟩; ext <;> simp
267265
map_add' _ _ := rfl
268266
map_smul' _ _ := rfl

0 commit comments

Comments
 (0)