Skip to content

Commit 7eeaaff

Browse files
committed
chore: classify porting notes about additional necessary beta reduction (#12130)
This subsumes some of the notes in #10752 and #10971. I'm on the fence as to whether replacing the dsimp only by beta_reduce is useful; this is easy to revert if needed.
1 parent 038cb58 commit 7eeaaff

File tree

16 files changed

+84
-80
lines changed

16 files changed

+84
-80
lines changed

Mathlib/Algebra/GCDMonoid/Basic.lean

Lines changed: 33 additions & 32 deletions
Original file line numberDiff line numberDiff line change
@@ -1048,15 +1048,15 @@ noncomputable def gcdMonoidOfGCD [DecidableEq α] (gcd : α → α → α)
10481048
lcm := fun a b =>
10491049
if a = 0 then 0 else Classical.choose ((gcd_dvd_left a b).trans (Dvd.intro b rfl))
10501050
gcd_mul_lcm := fun a b => by
1051-
-- Porting note (#10971): need `dsimp only` before `split_ifs`
1052-
dsimp only
1051+
-- Porting note(#12129): additional beta reduction needed
1052+
beta_reduce
10531053
split_ifs with a0
10541054
· rw [mul_zero, a0, zero_mul]
10551055
· rw [← Classical.choose_spec ((gcd_dvd_left a b).trans (Dvd.intro b rfl))]
10561056
lcm_zero_left := fun a => if_pos rfl
10571057
lcm_zero_right := fun a => by
1058-
-- Porting note (#10971): need `dsimp only` before `split_ifs`
1059-
dsimp only
1058+
-- Porting note(#12129): additional beta reduction needed
1059+
beta_reduce
10601060
split_ifs with a0
10611061
· rfl
10621062
have h := (Classical.choose_spec ((gcd_dvd_left a 0).trans (Dvd.intro 0 rfl))).symm
@@ -1108,17 +1108,17 @@ noncomputable def normalizedGCDMonoidOfGCD [NormalizationMonoid α] [DecidableEq
11081108
rw [← normalize_gcd] at this
11091109
rwa [normalize.map_mul, normalize_gcd, mul_right_inj' h1] at h2
11101110
gcd_mul_lcm := fun a b => by
1111-
-- Porting note (#10971): need `dsimp only`
1112-
dsimp only
1111+
-- Porting note(#12129): additional beta reduction needed
1112+
beta_reduce
11131113
split_ifs with a0
11141114
· rw [mul_zero, a0, zero_mul]
11151115
· rw [←
11161116
Classical.choose_spec (dvd_normalize_iff.2 ((gcd_dvd_left a b).trans (Dvd.intro b rfl)))]
11171117
exact normalize_associated (a * b)
11181118
lcm_zero_left := fun a => if_pos rfl
11191119
lcm_zero_right := fun a => by
1120-
-- Porting note (#10971): need `dsimp only`
1121-
dsimp only
1120+
-- Porting note(#12129): additional beta reduction needed
1121+
beta_reduce
11221122
split_ifs with a0
11231123
· rfl
11241124
rw [← normalize_eq_zero] at a0
@@ -1141,17 +1141,17 @@ noncomputable def gcdMonoidOfLCM [DecidableEq α] (lcm : α → α → α)
11411141
{ lcm
11421142
gcd := fun a b => if a = 0 then b else if b = 0 then a else Classical.choose (exists_gcd a b)
11431143
gcd_mul_lcm := fun a b => by
1144-
-- Porting note (#10971): need `dsimp only`
1145-
dsimp only
1144+
-- Porting note(#12129): additional beta reduction needed
1145+
beta_reduce
11461146
split_ifs with h h_1
11471147
· rw [h, eq_zero_of_zero_dvd (dvd_lcm_left _ _), mul_zero, zero_mul]
11481148
· 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 _ _)
11521152
gcd_dvd_left := fun a b => by
1153-
-- Porting note (#10971): need `dsimp only`
1154-
dsimp only
1153+
-- Porting note(#12129): additional beta reduction needed
1154+
beta_reduce
11551155
split_ifs with h h_1
11561156
· rw [h]
11571157
apply dvd_zero
@@ -1167,8 +1167,8 @@ noncomputable def gcdMonoidOfLCM [DecidableEq α] (lcm : α → α → α)
11671167
mul_dvd_mul_iff_right h]
11681168
apply dvd_lcm_right
11691169
gcd_dvd_right := fun a b => by
1170-
-- Porting note (#10971): need `dsimp only`
1171-
dsimp only
1170+
-- Porting note(#12129): additional beta reduction needed
1171+
beta_reduce
11721172
split_ifs with h h_1
11731173
· exact dvd_rfl
11741174
· rw [h_1]
@@ -1184,8 +1184,8 @@ noncomputable def gcdMonoidOfLCM [DecidableEq α] (lcm : α → α → α)
11841184
mul_dvd_mul_iff_right h_1]
11851185
apply dvd_lcm_left
11861186
dvd_gcd := fun {a b c} ac ab => by
1187-
-- Porting note (#10971): need `dsimp only`
1188-
dsimp only
1187+
-- Porting note(#12129): additional beta reduction needed
1188+
beta_reduce
11891189
split_ifs with h h_1
11901190
· exact ab
11911191
· exact ac
@@ -1219,7 +1219,7 @@ noncomputable def normalizedGCDMonoidOfLCM [NormalizationMonoid α] [DecidableEq
12191219
if a = 0 then normalize b
12201220
else if b = 0 then normalize a else Classical.choose (exists_gcd a b)
12211221
gcd_mul_lcm := fun a b => by
1222-
dsimp only
1222+
beta_reduce
12231223
split_ifs with h h_1
12241224
· rw [h, eq_zero_of_zero_dvd (dvd_lcm_left _ _), mul_zero, zero_mul]
12251225
· rw [h_1, eq_zero_of_zero_dvd (dvd_lcm_right _ _), mul_zero, mul_zero]
@@ -1247,7 +1247,7 @@ noncomputable def normalizedGCDMonoidOfLCM [NormalizationMonoid α] [DecidableEq
12471247
lcm_zero_left := fun a => eq_zero_of_zero_dvd (dvd_lcm_left _ _)
12481248
lcm_zero_right := fun a => eq_zero_of_zero_dvd (dvd_lcm_right _ _)
12491249
gcd_dvd_left := fun a b => by
1250-
dsimp only
1250+
beta_reduce
12511251
split_ifs with h h_1
12521252
· rw [h]
12531253
apply dvd_zero
@@ -1263,7 +1263,7 @@ noncomputable def normalizedGCDMonoidOfLCM [NormalizationMonoid α] [DecidableEq
12631263
mul_comm, mul_dvd_mul_iff_right h]
12641264
apply dvd_lcm_right
12651265
gcd_dvd_right := fun a b => by
1266-
dsimp only
1266+
beta_reduce
12671267
split_ifs with h h_1
12681268
· exact (normalize_associated _).dvd
12691269
· rw [h_1]
@@ -1279,7 +1279,7 @@ noncomputable def normalizedGCDMonoidOfLCM [NormalizationMonoid α] [DecidableEq
12791279
mul_dvd_mul_iff_right h_1]
12801280
apply dvd_lcm_left
12811281
dvd_gcd := fun {a b c} ac ab => by
1282-
dsimp only
1282+
beta_reduce
12831283
split_ifs with h h_1
12841284
· apply dvd_normalize_iff.2 ab
12851285
· apply dvd_normalize_iff.2 ac
@@ -1357,36 +1357,37 @@ instance (priority := 100) : NormalizedGCDMonoid G₀ where
13571357
normUnit x := if h : x = 0 then 1 else (Units.mk0 x h)⁻¹
13581358
normUnit_zero := dif_pos rfl
13591359
normUnit_mul := fun {x y} x0 y0 => Units.eq_iff.1 (by
1360-
-- Porting note (#10971): need `dsimp only`, also `simp` reaches maximum heartbeat
1360+
-- Porting note(#12129): additional beta reduction needed
1361+
-- Porting note: `simp` reaches maximum heartbeat
13611362
-- by Units.eq_iff.mp (by simp only [x0, y0, mul_comm])
1362-
dsimp only
1363+
beta_reduce
13631364
split_ifs with h
13641365
· rw [mul_eq_zero] at h
13651366
cases h
13661367
· exact absurd ‹x = 0› x0
13671368
· exact absurd ‹y = 0› y0
13681369
· rw [Units.mk0_mul, mul_inv_rev, mul_comm] )
13691370
normUnit_coe_units u := by
1370-
-- Porting note (#10971): need `dsimp only`
1371-
dsimp only
1371+
-- Porting note(#12129): additional beta reduction needed
1372+
beta_reduce
13721373
rw [dif_neg (Units.ne_zero _), Units.mk0_val]
13731374
gcd a b := if a = 0 ∧ b = 0 then 0 else 1
13741375
lcm a b := if a = 0 ∨ b = 0 then 0 else 1
13751376
gcd_dvd_left a b := by
1376-
-- Porting note (#10971): need `dsimp only`
1377-
dsimp only
1377+
-- Porting note(#12129): additional beta reduction needed
1378+
beta_reduce
13781379
split_ifs with h
13791380
· rw [h.1]
13801381
· exact one_dvd _
13811382
gcd_dvd_right a b := by
1382-
-- Porting note (#10971): need `dsimp only`
1383-
dsimp only
1383+
-- Porting note(#12129): additional beta reduction needed
1384+
beta_reduce
13841385
split_ifs with h
13851386
· rw [h.2]
13861387
· exact one_dvd _
13871388
dvd_gcd := fun {a b c} hac hab => by
1388-
-- Porting note (#10971): need `dsimp only`
1389-
dsimp only
1389+
-- Porting note(#12129): additional beta reduction needed
1390+
beta_reduce
13901391
split_ifs with h
13911392
· apply dvd_zero
13921393
· rw [not_and_or] at h
@@ -1402,8 +1403,8 @@ instance (priority := 100) : NormalizedGCDMonoid G₀ where
14021403
· by_cases hb : b = 0
14031404
· simp only [hb, and_true, or_true, ite_true, mul_zero]
14041405
exact Associated.refl _
1405-
-- Porting note (#10971): need `dsimp only`
1406-
· dsimp only
1406+
-- Porting note(#12129): additional beta reduction needed
1407+
· beta_reduce
14071408
rw [if_neg (not_and_of_not_left _ ha), one_mul, if_neg (not_or_of_not ha hb)]
14081409
exact (associated_one_iff_isUnit.mpr ((IsUnit.mk0 _ ha).mul (IsUnit.mk0 _ hb))).symm
14091410
lcm_zero_left b := if_pos (Or.inl rfl)

Mathlib/Algebra/GroupPower/IterateHom.lean

Lines changed: 3 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -175,12 +175,13 @@ section Semigroup
175175

176176
variable [Semigroup G] {a b c : G}
177177

178-
-- Porting note (#10971): need `dsimp only`, see https://leanprover.zulipchat.com/#narrow/stream/
178+
-- Porting note(#12129): additional beta reduction needed
179+
-- see also https://leanprover.zulipchat.com/#narrow/stream/
179180
-- 287929-mathlib4/topic/dsimp.20before.20rw/near/317063489
180181
@[to_additive]
181182
theorem SemiconjBy.function_semiconj_mul_left (h : SemiconjBy a b c) :
182183
Function.Semiconj (a * ·) (b * ·) (c * ·) := fun j => by
183-
dsimp only; rw [← mul_assoc, h.eq, mul_assoc]
184+
beta_reduce; rw [← mul_assoc, h.eq, mul_assoc]
184185
#align semiconj_by.function_semiconj_mul_left SemiconjBy.function_semiconj_mul_left
185186
#align add_semiconj_by.function_semiconj_add_left AddSemiconjBy.function_semiconj_add_left
186187

Mathlib/Algebra/MonoidAlgebra/Basic.lean

Lines changed: 8 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -699,8 +699,8 @@ def liftMagma [Module k A] [IsScalarTower k A A] [SMulCommClass k A A] :
699699
{ liftAddHom fun x => (smulAddHom k A).flip (f x) with
700700
toFun := fun a => a.sum fun m t => t • f m
701701
map_smul' := fun t' a => by
702-
-- Porting note: `dsimp` is required for beta reduction.
703-
dsimp only []
702+
-- Porting note(#12129): additional beta reduction needed
703+
beta_reduce
704704
rw [Finsupp.smul_sum, sum_smul_index']
705705
· simp_rw [smul_assoc, MonoidHom.id_apply]
706706
· intro m
@@ -1038,8 +1038,8 @@ def equivariantOfLinearOfComm : V →ₗ[MonoidAlgebra k G] W where
10381038
toFun := f
10391039
map_add' v v' := by simp
10401040
map_smul' c v := by
1041-
-- Porting note: `dsimp` is required for beta reduction.
1042-
dsimp only []
1041+
-- Porting note(#12129): additional beta reduction needed
1042+
beta_reduce
10431043
-- Porting note: Was `apply`.
10441044
refine Finsupp.induction c ?_ ?_
10451045
· simp
@@ -1757,8 +1757,8 @@ protected def AddMonoidAlgebra.toMultiplicative [Semiring k] [Add G] :
17571757
Multiplicative.ofAdd with
17581758
toFun := equivMapDomain Multiplicative.ofAdd
17591759
map_mul' := fun x y => by
1760-
-- Porting note: `dsimp` is required for beta reduction.
1761-
dsimp only []
1760+
-- Porting note: added `dsimp only`; `beta_reduce` alone is not sufficient
1761+
dsimp only
17621762
repeat' rw [equivMapDomain_eq_mapDomain (M := k)]
17631763
dsimp [Multiplicative.ofAdd]
17641764
exact MonoidAlgebra.mapDomain_mul (α := Multiplicative G) (β := k)
@@ -1771,8 +1771,8 @@ protected def MonoidAlgebra.toAdditive [Semiring k] [Mul G] :
17711771
{ Finsupp.domCongr Additive.ofMul with
17721772
toFun := equivMapDomain Additive.ofMul
17731773
map_mul' := fun x y => by
1774-
-- Porting note: `dsimp` is required for beta reduction.
1775-
dsimp only []
1774+
-- Porting note: added `dsimp only`; `beta_reduce` alone is not sufficient
1775+
dsimp only
17761776
repeat' rw [equivMapDomain_eq_mapDomain (M := k)]
17771777
dsimp [Additive.ofMul]
17781778
convert MonoidAlgebra.mapDomain_mul (β := k) (MulHom.id G) x y }

Mathlib/Algebra/Polynomial/Coeff.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -81,8 +81,8 @@ def lsum {R A M : Type*} [Semiring R] [Semiring A] [AddCommMonoid M] [Module R A
8181
toFun p := p.sum (f · ·)
8282
map_add' p q := sum_add_index p q _ (fun n => (f n).map_zero) fun n _ _ => (f n).map_add _ _
8383
map_smul' c p := by
84-
-- Porting note: `dsimp only []` is required for beta reduction.
85-
dsimp only []
84+
-- Porting note: added `dsimp only`; `beta_reduce` alone is not sufficient
85+
dsimp only
8686
rw [sum_eq_of_subset (f · ·) (fun n => (f n).map_zero) (support_smul c p)]
8787
simp only [sum_def, Finset.smul_sum, coeff_smul, LinearMap.map_smul, RingHom.id_apply]
8888
#align polynomial.lsum Polynomial.lsum

Mathlib/Computability/TuringMachine.lean

Lines changed: 2 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -2688,11 +2688,10 @@ theorem tr_respects_aux {q v T k} {S : ∀ k, List (Γ k)}
26882688
attribute [local simp] Respects TM2.step TM2.stepAux trNormal
26892689

26902690
theorem tr_respects : Respects (TM2.step M) (TM1.step (tr M)) TrCfg := by
2691-
-- Porting note: `simp only`s are required for beta reductions.
2691+
-- Porting note(#12129): additional beta reduction needed
26922692
intro c₁ c₂ h
26932693
cases' h with l v S L hT
26942694
cases' l with l; · constructor
2695-
simp only [TM2.step, Respects, Option.map_some']
26962695
rsuffices ⟨b, c, r⟩ : ∃ b, _ ∧ Reaches (TM1.step (tr M)) _ _
26972696
· exact ⟨b, c, TransGen.head' rfl r⟩
26982697
simp only [tr]
@@ -2703,7 +2702,7 @@ theorem tr_respects : Respects (TM2.step M) (TM1.step (tr M)) TrCfg := by
27032702
| H₂ a _ IH => exact IH _ hT
27042703
| H₃ p q₁ q₂ IH₁ IH₂ =>
27052704
unfold TM2.stepAux trNormal TM1.stepAux
2706-
simp only []
2705+
beta_reduce
27072706
cases p v <;> [exact IH₂ _ hT; exact IH₁ _ hT]
27082707
| H₄ => exact ⟨_, ⟨_, hT⟩, ReflTransGen.refl⟩
27092708
| H₅ => exact ⟨_, ⟨_, hT⟩, ReflTransGen.refl⟩

Mathlib/Data/Finsupp/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -881,7 +881,7 @@ def filter (p : α → Prop) [DecidablePred p] (f : α →₀ M) : α →₀ M w
881881
toFun a := if p a then f a else 0
882882
support := f.support.filter p
883883
mem_support_toFun a := by
884-
simp only -- Porting note: necessary to beta reduce to activate `split_ifs`
884+
beta_reduce -- Porting note(#12129): additional beta reduction needed to activate `split_ifs`
885885
split_ifs with h <;>
886886
· simp only [h, mem_filter, mem_support_iff]
887887
tauto

Mathlib/FieldTheory/RatFunc.lean

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -607,7 +607,7 @@ def map [MonoidHomClass F R[X] S[X]] (φ : F) (hφ : R[X]⁰ ≤ S[X]⁰.comap
607607
RatFunc.liftOn f
608608
(fun n d => if h : φ d ∈ S[X]⁰ then ofFractionRing (Localization.mk (φ n) ⟨φ d, h⟩) else 0)
609609
fun {p q p' q'} hq hq' h => by
610-
dsimp only -- Porting note: force the function to be applied
610+
beta_reduce -- Porting note(#12129): force the function to be applied
611611
rw [dif_pos, dif_pos]
612612
congr 1 -- Porting note: this was a `rw [ofFractionRing.inj_eq]` which was overkill anyway
613613
rw [Localization.mk_eq_mk_iff]
@@ -617,12 +617,12 @@ def map [MonoidHomClass F R[X] S[X]] (φ : F) (hφ : R[X]⁰ ≤ S[X]⁰.comap
617617
refine' Localization.r_of_eq _
618618
simpa only [map_mul] using congr_arg φ h
619619
map_one' := by
620-
dsimp only -- Porting note: force the function to be applied
620+
beta_reduce -- Porting note(#12129): force the function to be applied
621621
rw [← ofFractionRing_one, ← Localization.mk_one, liftOn_ofFractionRing_mk, dif_pos]
622622
· simpa using ofFractionRing_one
623623
· simpa using Submonoid.one_mem _
624624
map_mul' x y := by
625-
dsimp only -- Porting note: force the function to be applied
625+
beta_reduce -- Porting note(#12129): force the function to be applied
626626
cases' x with x; cases' y with y
627627
-- Porting note: added `using Localization.rec` (`Localization.induction_on` didn't work)
628628
induction' x using Localization.rec with p q
@@ -699,7 +699,7 @@ def liftMonoidWithZeroHom (φ : R[X] →*₀ G₀) (hφ : R[X]⁰ ≤ G₀⁰.co
699699
rw [div_eq_div_iff, ← map_mul, mul_comm p, h, map_mul, mul_comm] <;>
700700
exact nonZeroDivisors.ne_zero (hφ ‹_›)
701701
map_one' := by
702-
dsimp only -- Porting note: force the function to be applied
702+
dsimp only -- Porting note: force the function to be applied (not just beta reduction!)
703703
rw [← ofFractionRing_one, ← Localization.mk_one, liftOn_ofFractionRing_mk]
704704
simp only [map_one, OneMemClass.coe_one, div_one]
705705
map_mul' x y := by
@@ -712,7 +712,7 @@ def liftMonoidWithZeroHom (φ : R[X] →*₀ G₀) (hφ : R[X]⁰ ≤ G₀⁰.co
712712
· rfl
713713
· rfl
714714
map_zero' := by
715-
dsimp only -- Porting note: force the function to be applied
715+
beta_reduce -- Porting note(#12129): force the function to be applied
716716
rw [← ofFractionRing_zero, ← Localization.mk_zero (1 : R[X]⁰), liftOn_ofFractionRing_mk]
717717
simp only [map_zero, zero_div]
718718
#align ratfunc.lift_monoid_with_zero_hom RatFunc.liftMonoidWithZeroHom

Mathlib/GroupTheory/OrderOfElement.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -42,10 +42,10 @@ variable [Monoid G] {a b x y : G} {n m : ℕ}
4242

4343
section IsOfFinOrder
4444

45-
-- Porting note: we need a `dsimp` in the middle of the rewrite to do beta reduction
45+
-- Porting note(#12129): additional beta reduction needed
4646
@[to_additive]
4747
theorem isPeriodicPt_mul_iff_pow_eq_one (x : G) : IsPeriodicPt (x * ·) n 1 ↔ x ^ n = 1 := by
48-
rw [IsPeriodicPt, IsFixedPt, mul_left_iterate]; dsimp; rw [mul_one]
48+
rw [IsPeriodicPt, IsFixedPt, mul_left_iterate]; beta_reduce; rw [mul_one]
4949
#align is_periodic_pt_mul_iff_pow_eq_one isPeriodicPt_mul_iff_pow_eq_one
5050
#align is_periodic_pt_add_iff_nsmul_eq_zero isPeriodicPt_add_iff_nsmul_eq_zero
5151

@@ -176,8 +176,8 @@ protected lemma IsOfFinOrder.orderOf_pos (h : IsOfFinOrder x) : 0 < orderOf x :=
176176
theorem pow_orderOf_eq_one (x : G) : x ^ orderOf x = 1 := by
177177
-- Porting note: was `convert`, but the `1` in the lemma is equal only after unfolding
178178
refine Eq.trans ?_ (isPeriodicPt_minimalPeriod (x * ·) 1)
179-
-- Porting note: we need a `dsimp` in the middle of the rewrite to do beta reduction
180-
rw [orderOf, mul_left_iterate]; dsimp; rw [mul_one]
179+
-- Porting note(#12129): additional beta reduction needed in the middle of the rewrite
180+
rw [orderOf, mul_left_iterate]; beta_reduce; rw [mul_one]
181181
#align pow_order_of_eq_one pow_orderOf_eq_one
182182
#align add_order_of_nsmul_eq_zero addOrderOf_nsmul_eq_zero
183183

Mathlib/GroupTheory/Transfer.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -100,7 +100,9 @@ the transfer homomorphism is `transfer ϕ : G →+ A`."]
100100
noncomputable def transfer [FiniteIndex H] : G →* A :=
101101
let T : leftTransversals (H : Set G) := Inhabited.default
102102
{ toFun := fun g => diff ϕ T (g • T)
103-
map_one' := by simp only; rw [one_smul, diff_self] -- Porting note: added `simp only`
103+
-- Porting note(#12129): additional beta reduction needed
104+
map_one' := by beta_reduce; rw [one_smul, diff_self]
105+
-- Porting note: added `simp only` (not just beta reduction)
104106
map_mul' := fun g h => by simp only; rw [mul_smul, ← diff_mul_diff, smul_diff_smul] }
105107
#align monoid_hom.transfer MonoidHom.transfer
106108
#align add_monoid_hom.transfer AddMonoidHom.transfer

Mathlib/LinearAlgebra/Basis.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -313,8 +313,8 @@ theorem repr_apply_eq (f : M → ι → R) (hadd : ∀ x y, f (x + y) = f x + f
313313
(x : M) (i : ι) : b.repr x i = f x i := by
314314
let f_i : M →ₗ[R] R :=
315315
{ toFun := fun x => f x i
316-
-- Porting note: `dsimp only []` is required for beta reduction.
317-
map_add' := fun _ _ => by dsimp only []; rw [hadd, Pi.add_apply]
316+
-- Porting note(#12129): additional beta reduction needed
317+
map_add' := fun _ _ => by beta_reduce; rw [hadd, Pi.add_apply]
318318
map_smul' := fun _ _ => by simp [hsmul, Pi.smul_apply] }
319319
have : Finsupp.lapply i ∘ₗ ↑b.repr = f_i := by
320320
refine' b.ext fun j => _

0 commit comments

Comments
 (0)