Skip to content

Commit

Permalink
feat(Algebra/Module): Use coercion from SemilinearEquivClass to Semil…
Browse files Browse the repository at this point in the history
…inearEquiv (#11966)
  • Loading branch information
Amelia Livingston authored and callesonne committed Apr 22, 2024
1 parent 5f92bfe commit 88e7307
Showing 1 changed file with 13 additions and 0 deletions.
13 changes: 13 additions & 0 deletions Mathlib/Algebra/Module/Equiv.lean
Original file line number Diff line number Diff line change
Expand Up @@ -122,6 +122,19 @@ instance (priority := 100) [RingHomInvPair σ σ'] [RingHomInvPair σ' σ]
[EquivLike F M M₂] [s : SemilinearEquivClass F σ M M₂] : SemilinearMapClass F σ M M₂ :=
{ s with }

variable {F}

/-- Reinterpret an element of a type of semilinear equivalences as a semilinear equivalence. -/
@[coe]
def semilinearEquiv [RingHomInvPair σ σ'] [RingHomInvPair σ' σ]
[EquivLike F M M₂] [SemilinearEquivClass F σ M M₂] (f : F) : M ≃ₛₗ[σ] M₂ :=
{ (f : M ≃+ M₂), (f : M →ₛₗ[σ] M₂) with }

/-- Reinterpret an element of a type of semilinear equivalences as a semilinear equivalence. -/
instance instCoeToSemilinearEquiv [RingHomInvPair σ σ'] [RingHomInvPair σ' σ]
[EquivLike F M M₂] [SemilinearEquivClass F σ M M₂] : CoeHead F (M ≃ₛₗ[σ] M₂) where
coe f := semilinearEquiv f

end SemilinearEquivClass

namespace LinearEquiv
Expand Down

0 comments on commit 88e7307

Please sign in to comment.