Skip to content

Commit

Permalink
refactor: fix LinearMap.coe_mk (#2336)
Browse files Browse the repository at this point in the history
Co-authored-by: Pol_tta <52843868+Komyyy@users.noreply.github.com>
  • Loading branch information
Komyyy and Komyyy committed Feb 18, 2023
1 parent 4557f2c commit 79902bf
Showing 1 changed file with 7 additions and 1 deletion.
8 changes: 7 additions & 1 deletion Mathlib/Algebra/Module/LinearMap.lean
Expand Up @@ -269,11 +269,17 @@ protected def Simps.apply {R S : Type _} [Semiring R] [Semiring S] (σ : R →+*
initialize_simps_projections LinearMap (toAddHom_toFun → apply)

@[simp]
theorem coe_mk {σ : R →+* S} (f : M →+ M₃) (h) :
theorem coe_mk {σ : R →+* S} (f : AddHom M M₃) (h) :
((LinearMap.mk f h : M →ₛₗ[σ] M₃) : M → M₃) = f :=
rfl
#align linear_map.coe_mk LinearMap.coe_mk

-- Porting note: This theorem is new.
@[simp]
theorem coe_addHom_mk {σ : R →+* S} (f : AddHom M M₃) (h) :
((LinearMap.mk f h : M →ₛₗ[σ] M₃) : AddHom M M₃) = f :=
rfl

/-- Identity map as a `LinearMap` -/
def id : M →ₗ[R] M :=
{ DistribMulActionHom.id R with toFun := _root_.id }
Expand Down

0 comments on commit 79902bf

Please sign in to comment.