diff --git a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean index 845c92e62905f..de9a7c0377ebb 100644 --- a/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Subalgebra/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Yury Kudryashov ! This file was ported from Lean 3 source module algebra.algebra.subalgebra.basic -! leanprover-community/mathlib commit c4658a649d216f57e99621708b09dcb3dcccbd23 +! leanprover-community/mathlib commit 8130e5155d637db35907c272de9aec9dc851c03a ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -123,6 +123,8 @@ theorem smul_mem {x : A} (hx : x ∈ S) (r : R) : r • x ∈ S := (Algebra.smul_def r x).symm ▸ mul_mem (S.algebraMap_mem r) hx #align subalgebra.smul_mem Subalgebra.smul_mem +instance : SMulMemClass (Subalgebra R A) R A where smul_mem r _x hx := smul_mem _ hx r + protected theorem one_mem : (1 : A) ∈ S := one_mem S #align subalgebra.one_mem Subalgebra.one_mem diff --git a/Mathlib/Algebra/Module/Submodule/Basic.lean b/Mathlib/Algebra/Module/Submodule/Basic.lean index b4885f293bc80..d9c9757bcfebb 100644 --- a/Mathlib/Algebra/Module/Submodule/Basic.lean +++ b/Mathlib/Algebra/Module/Submodule/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro ! This file was ported from Lean 3 source module algebra.module.submodule.basic -! leanprover-community/mathlib commit 155d5519569cecf21f48c534da8b729890e20ce6 +! leanprover-community/mathlib commit 8130e5155d637db35907c272de9aec9dc851c03a ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -38,15 +38,6 @@ universe u'' u' u v w variable {G : Type u''} {S : Type u'} {R : Type u} {M : Type v} {ι : Type w} -/-- -`SubmoduleClass S R M` says `S` is a type of submodules `s ≤ M`. - -Note that only `R` is marked as `outParam` since `M` is already supplied by the `SetLike` class. --/ -class SubmoduleClass (S : Type _) (R : outParam <| Type _) (M : Type _) [AddZeroClass M] [SMul R M] - [SetLike S M] [AddSubmonoidClass S M] extends SMulMemClass S R M -#align submodule_class SubmoduleClass - /-- A submodule of a module is one which is closed under vector operations. This is a sufficient condition for the subset of vectors in the submodule to themselves form a module. -/ @@ -76,9 +67,9 @@ instance addSubmonoidClass : AddSubmonoidClass (Submodule R M) M where add_mem := AddSubsemigroup.add_mem' _ #align submodule.add_submonoid_class Submodule.addSubmonoidClass -instance submoduleClass : SubmoduleClass (Submodule R M) R M where +instance smulMemClass : SMulMemClass (Submodule R M) R M where smul_mem {s} c _ h := SubMulAction.smul_mem' s.toSubMulAction c h -#align submodule.submodule_class Submodule.submoduleClass +#align submodule.smul_mem_class Submodule.smulMemClass @[simp] theorem mem_toAddSubmonoid (p : Submodule R M) (x : M) : x ∈ p.toAddSubmonoid ↔ x ∈ p := @@ -185,30 +176,30 @@ theorem coe_toSubMulAction (p : Submodule R M) : (p.toSubMulAction : Set M) = p end Submodule -namespace SubmoduleClass +namespace SMulMemClass variable [Semiring R] [AddCommMonoid M] [Module R M] {A : Type _} [SetLike A M] - [AddSubmonoidClass A M] [SubmoduleClass A R M] (S' : A) + [AddSubmonoidClass A M] [SMulMemClass A R M] (S' : A) --- Prefer subclasses of `Module` over `SubmoduleClass`. +-- Prefer subclasses of `Module` over `SMulMemClass`. /-- A submodule of a `Module` is a `Module`. -/ instance (priority := 75) toModule : Module R S' := Subtype.coe_injective.module R (AddSubmonoidClass.Subtype S') (SetLike.val_smul S') -#align submodule_class.to_module SubmoduleClass.toModule +#align submodule_class.to_module SMulMemClass.toModule /-- The natural `R`-linear map from a submodule of an `R`-module `M` to `M`. -/ protected def subtype : S' →ₗ[R] M where toFun := Subtype.val map_add' _ _ := rfl map_smul' _ _ := rfl -#align submodule_class.subtype SubmoduleClass.subtype +#align submodule_class.subtype SMulMemClass.subtype @[simp] -protected theorem coeSubtype : (SubmoduleClass.subtype S' : S' → M) = Subtype.val := +protected theorem coeSubtype : (SMulMemClass.subtype S' : S' → M) = Subtype.val := rfl -#align submodule_class.coe_subtype SubmoduleClass.coeSubtype +#align submodule_class.coe_subtype SMulMemClass.coeSubtype -end SubmoduleClass +end SMulMemClass namespace Submodule