Skip to content

Commit

Permalink
fix: resolve some nolint simpNF commands (#7929)
Browse files Browse the repository at this point in the history
  • Loading branch information
fpvandoorn committed Oct 27, 2023
1 parent f52e271 commit 220023c
Show file tree
Hide file tree
Showing 4 changed files with 3 additions and 14 deletions.
5 changes: 1 addition & 4 deletions Mathlib/Algebra/Algebra/Subalgebra/Basic.lean
Expand Up @@ -743,16 +743,13 @@ noncomputable def ofInjectiveField {E F : Type*} [DivisionRing E] [Semiring F] [
/-- Given an equivalence `e : A ≃ₐ[R] B` of `R`-algebras and a subalgebra `S` of `A`,
`subalgebra_map` is the induced equivalence between `S` and `S.map e` -/
@[simps!]
def subalgebraMap (e : A ≃ₐ[R] B) (S : Subalgebra R A) : S ≃ₐ[R] S.map e.toAlgHom :=
def subalgebraMap (e : A ≃ₐ[R] B) (S : Subalgebra R A) : S ≃ₐ[R] S.map (e : A →ₐ[R] B) :=
{ e.toRingEquiv.subsemiringMap S.toSubsemiring with
commutes' := fun r => by
ext; dsimp only; erw [RingEquiv.subsemiringMap_apply_coe]
exact e.commutes _ }
#align alg_equiv.subalgebra_map AlgEquiv.subalgebraMap

-- These lemmas have always been bad (#7657), but lean4#2644 made `simp` start noticing
attribute [nolint simpNF] AlgEquiv.subalgebraMap_apply_coe AlgEquiv.subalgebraMap_symm_apply_coe

end AlgEquiv

namespace Algebra
Expand Down
5 changes: 1 addition & 4 deletions Mathlib/Algebra/Category/GroupCat/Basic.lean
Expand Up @@ -145,7 +145,7 @@ set_option linter.uppercaseLean3 false in
/-- Typecheck an `AddMonoidHom` as a morphism in `AddGroup`. -/
add_decl_doc AddGroupCat.ofHom

@[to_additive (attr := simp)]
@[to_additive]
theorem ofHom_apply {X Y : Type _} [Group X] [Group Y] (f : X →* Y) (x : X) :
(ofHom f) x = f x :=
rfl
Expand All @@ -154,9 +154,6 @@ set_option linter.uppercaseLean3 false in
set_option linter.uppercaseLean3 false in
#align AddGroup.of_hom_apply AddGroupCat.ofHom_apply

-- These lemmas have always been bad (#7657), but lean4#2644 made `simp` start noticing
attribute [nolint simpNF] AddGroupCat.ofHom_apply GroupCat.ofHom_apply

@[to_additive]
instance ofUnique (G : Type*) [Group G] [i : Unique G] : Unique (GroupCat.of G) := i
set_option linter.uppercaseLean3 false in
Expand Down
2 changes: 0 additions & 2 deletions Mathlib/Algebra/Category/ModuleCat/Monoidal/Closed.lean
Expand Up @@ -66,8 +66,6 @@ open MonoidalCategory

-- porting note: `CoeFun` was replaced by `FunLike`
-- I can't seem to express the function coercion here without writing `@FunLike.coe`.
-- These lemmas have always been bad (#7657), but lean4#2644 made `simp` start noticing
@[simp, nolint simpNF]
theorem monoidalClosed_curry {M N P : ModuleCat.{u} R} (f : M ⊗ N ⟶ P) (x : M) (y : N) :
@FunLike.coe _ _ _ LinearMap.instFunLike
((MonoidalClosed.curry f : N →ₗ[R] M →ₗ[R] P) y) x = f (x ⊗ₜ[R] y) :=
Expand Down
5 changes: 1 addition & 4 deletions Mathlib/LinearAlgebra/Finrank.lean
Expand Up @@ -497,15 +497,12 @@ theorem coe_basisOfTopLeSpanOfCardEqFinrank {ι : Type*} [Fintype ι] (b : ι
/-- A finset of `finrank K V` vectors forms a basis if they span the whole space. -/
@[simps! repr_apply]
noncomputable def finsetBasisOfTopLeSpanOfCardEqFinrank {s : Finset V}
(le_span : ⊤ ≤ span K (s : Set V)) (card_eq : s.card = finrank K V) : Basis (s : Set V) K V :=
(le_span : ⊤ ≤ span K (s : Set V)) (card_eq : s.card = finrank K V) : Basis {x // x ∈ s} K V :=
basisOfTopLeSpanOfCardEqFinrank ((↑) : ↥(s : Set V) → V)
((@Subtype.range_coe_subtype _ fun x => x ∈ s).symm ▸ le_span)
(_root_.trans (Fintype.card_coe _) card_eq)
#align finset_basis_of_top_le_span_of_card_eq_finrank finsetBasisOfTopLeSpanOfCardEqFinrank

-- These lemmas have always been bad (#7657), but lean4#2644 made `simp` start noticing
attribute [nolint simpNF] finsetBasisOfTopLeSpanOfCardEqFinrank_repr_apply

/-- A set of `finrank K V` vectors forms a basis if they span the whole space. -/
@[simps! repr_apply]
noncomputable def setBasisOfTopLeSpanOfCardEqFinrank {s : Set V} [Fintype s]
Expand Down

0 comments on commit 220023c

Please sign in to comment.