Skip to content

Commit

Permalink
chore: redistribute a few results from Algebra.Module.Submodule.Linea…
Browse files Browse the repository at this point in the history
…rMap (#12295)

These don't involve submodules at all.
  • Loading branch information
Ruben-VandeVelde committed Apr 22, 2024
1 parent f4476bd commit 5b04cac
Show file tree
Hide file tree
Showing 3 changed files with 29 additions and 29 deletions.
25 changes: 25 additions & 0 deletions Mathlib/Algebra/Module/LinearMap/Basic.lean
Expand Up @@ -934,6 +934,15 @@ theorem default_def : (default : M →ₛₗ[σ₁₂] M₂) = 0 :=
rfl
#align linear_map.default_def LinearMap.default_def

instance uniqueOfLeft [Subsingleton M] : Unique (M →ₛₗ[σ₁₂] M₂) :=
{ inferInstanceAs (Inhabited (M →ₛₗ[σ₁₂] M₂)) with
uniq := fun f => ext fun x => by rw [Subsingleton.elim x 0, map_zero, map_zero] }
#align linear_map.unique_of_left LinearMap.uniqueOfLeft

instance uniqueOfRight [Subsingleton M₂] : Unique (M →ₛₗ[σ₁₂] M₂) :=
coe_injective.unique
#align linear_map.unique_of_right LinearMap.uniqueOfRight

/-- The sum of two linear maps is linear. -/
instance : Add (M →ₛₗ[σ₁₂] M₂) :=
fun f g ↦
Expand Down Expand Up @@ -1009,6 +1018,22 @@ instance addCommGroup : AddCommGroup (M →ₛₗ[σ₁₂] N₂) :=
DFunLike.coe_injective.addCommGroup _ rfl (fun _ _ ↦ rfl) (fun _ ↦ rfl) (fun _ _ ↦ rfl)
(fun _ _ ↦ rfl) fun _ _ ↦ rfl

/-- Evaluation of a `σ₁₂`-linear map at a fixed `a`, as an `AddMonoidHom`. -/
@[simps]
def evalAddMonoidHom (a : M) : (M →ₛₗ[σ₁₂] M₂) →+ M₂ where
toFun f := f a
map_add' f g := LinearMap.add_apply f g a
map_zero' := rfl
#align linear_map.eval_add_monoid_hom LinearMap.evalAddMonoidHom

/-- `LinearMap.toAddMonoidHom` promoted to an `AddMonoidHom`. -/
@[simps]
def toAddMonoidHom' : (M →ₛₗ[σ₁₂] M₂) →+ M →+ M₂ where
toFun := toAddMonoidHom
map_zero' := by ext; rfl
map_add' := by intros; ext; rfl
#align linear_map.to_add_monoid_hom' LinearMap.toAddMonoidHom'

end Arithmetic

section Actions
Expand Down
4 changes: 4 additions & 0 deletions Mathlib/Algebra/Module/LinearMap/End.lean
Expand Up @@ -67,6 +67,10 @@ theorem coe_one : ⇑(1 : Module.End R M) = _root_.id := rfl
theorem coe_mul (f g : Module.End R M) : ⇑(f * g) = f ∘ g := rfl
#align linear_map.coe_mul LinearMap.coe_mul

instance _root_.Module.End.instNontrivial [Nontrivial M] : Nontrivial (Module.End R M) := by
obtain ⟨m, ne⟩ := exists_ne (0 : M)
exact nontrivial_of_ne 1 0 fun p => ne (LinearMap.congr_fun p m)

instance _root_.Module.End.monoid : Monoid (Module.End R M) where
mul := (· * ·)
one := (1 : M →ₗ[R] M)
Expand Down
29 changes: 0 additions & 29 deletions Mathlib/Algebra/Module/Submodule/LinearMap.lean
Expand Up @@ -211,40 +211,11 @@ theorem restrict_eq_domRestrict_codRestrict {f : M →ₗ[R] M₁} {p : Submodul
rfl
#align linear_map.restrict_eq_dom_restrict_cod_restrict LinearMap.restrict_eq_domRestrict_codRestrict

instance uniqueOfLeft [Subsingleton M] : Unique (M →ₛₗ[σ₁₂] M₂) :=
{ inferInstanceAs (Inhabited (M →ₛₗ[σ₁₂] M₂)) with
uniq := fun f => ext fun x => by rw [Subsingleton.elim x 0, map_zero, map_zero] }
#align linear_map.unique_of_left LinearMap.uniqueOfLeft

instance uniqueOfRight [Subsingleton M₂] : Unique (M →ₛₗ[σ₁₂] M₂) :=
coe_injective.unique
#align linear_map.unique_of_right LinearMap.uniqueOfRight

/-- Evaluation of a `σ₁₂`-linear map at a fixed `a`, as an `AddMonoidHom`. -/
@[simps]
def evalAddMonoidHom (a : M) : (M →ₛₗ[σ₁₂] M₂) →+ M₂ where
toFun f := f a
map_add' f g := LinearMap.add_apply f g a
map_zero' := rfl
#align linear_map.eval_add_monoid_hom LinearMap.evalAddMonoidHom

/-- `LinearMap.toAddMonoidHom` promoted to an `AddMonoidHom`. -/
@[simps]
def toAddMonoidHom' : (M →ₛₗ[σ₁₂] M₂) →+ M →+ M₂ where
toFun := toAddMonoidHom
map_zero' := by ext; rfl
map_add' := by intros; ext; rfl
#align linear_map.to_add_monoid_hom' LinearMap.toAddMonoidHom'

theorem sum_apply (t : Finset ι) (f : ι → M →ₛₗ[σ₁₂] M₂) (b : M) :
(∑ d in t, f d) b = ∑ d in t, f d b :=
_root_.map_sum ((AddMonoidHom.eval b).comp toAddMonoidHom') f _
#align linear_map.sum_apply LinearMap.sum_apply

instance [Nontrivial M] : Nontrivial (Module.End R M) := by
obtain ⟨m, ne⟩ := exists_ne (0 : M)
exact nontrivial_of_ne 1 0 fun p => ne (LinearMap.congr_fun p m)

@[simp, norm_cast]
theorem coeFn_sum {ι : Type*} (t : Finset ι) (f : ι → M →ₛₗ[σ₁₂] M₂) :
⇑(∑ i in t, f i) = ∑ i in t, (f i : M → M₂) :=
Expand Down

0 comments on commit 5b04cac

Please sign in to comment.