Documentation

Mathlib.Algebra.Algebra.Subalgebra.Lattice

Complete lattice structure of subalgebras #

In this file we define Algebra.adjoin and the complete lattice structure on subalgebras.

More lemmas about adjoin can be found in Mathlib/RingTheory/Adjoin/Basic.lean.

def Algebra.adjoin (R : Type u) {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] (s : Set A) :

The minimal subalgebra that includes s.

Equations
    Instances For

      Galois insertion between adjoin and coe.

      Equations
        Instances For
          theorem Algebra.sup_def {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] (S T : Subalgebra R A) :
          ST = adjoin R (S T)
          theorem Algebra.sSup_def {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] (S : Set (Subalgebra R A)) :
          @[simp]
          theorem Algebra.coe_top {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] :
          @[simp]
          theorem Algebra.mem_top {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {x : A} :
          @[simp]
          theorem Algebra.top_toSubring {R : Type u_1} {A : Type u_2} [CommRing R] [Ring A] [Algebra R A] :
          @[simp]
          theorem Algebra.toSubsemiring_eq_top {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {S : Subalgebra R A} :
          @[simp]
          theorem Algebra.toSubring_eq_top {R : Type u_1} {A : Type u_2} [CommRing R] [Ring A] [Algebra R A] {S : Subalgebra R A} :
          theorem Algebra.mem_sup_left {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {S T : Subalgebra R A} {x : A} :
          x Sx ST
          theorem Algebra.mem_sup_right {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {S T : Subalgebra R A} {x : A} :
          x Tx ST
          theorem Algebra.mul_mem_sup {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {S T : Subalgebra R A} {x y : A} (hx : x S) (hy : y T) :
          x * y ST
          theorem Algebra.map_sup {R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A →ₐ[R] B) (S T : Subalgebra R A) :
          theorem Algebra.map_inf {R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A →ₐ[R] B) (hf : Function.Injective f) (S T : Subalgebra R A) :
          @[simp]
          theorem Algebra.coe_inf {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] (S T : Subalgebra R A) :
          (ST) = S T
          @[simp]
          theorem Algebra.mem_inf {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {S T : Subalgebra R A} {x : A} :
          x ST x S x T
          @[simp]
          theorem Algebra.inf_toSubsemiring {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] (S T : Subalgebra R A) :
          @[simp]
          theorem Algebra.sup_toSubsemiring {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] (S T : Subalgebra R A) :
          @[simp]
          theorem Algebra.coe_sInf {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] (S : Set (Subalgebra R A)) :
          (sInf S) = sS, s
          theorem Algebra.mem_sInf {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {S : Set (Subalgebra R A)} {x : A} :
          x sInf S pS, x p
          @[simp]
          theorem Algebra.coe_iInf {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {ι : Sort u_1} {S : ιSubalgebra R A} :
          (⨅ (i : ι), S i) = ⋂ (i : ι), (S i)
          theorem Algebra.mem_iInf {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {ι : Sort u_1} {S : ιSubalgebra R A} {x : A} :
          x ⨅ (i : ι), S i ∀ (i : ι), x S i
          theorem Algebra.map_iInf {R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {ι : Sort u_1} [Nonempty ι] (f : A →ₐ[R] B) (hf : Function.Injective f) (s : ιSubalgebra R A) :
          Subalgebra.map f (iInf s) = ⨅ (i : ι), Subalgebra.map f (s i)
          @[simp]
          theorem Algebra.iInf_toSubmodule {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {ι : Sort u_1} (S : ιSubalgebra R A) :
          Subalgebra.toSubmodule (⨅ (i : ι), S i) = ⨅ (i : ι), Subalgebra.toSubmodule (S i)
          @[simp]
          theorem Algebra.iInf_toSubsemiring {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {ι : Sort u_1} (S : ιSubalgebra R A) :
          (iInf S).toSubsemiring = ⨅ (i : ι), (S i).toSubsemiring
          @[simp]
          theorem Algebra.iSup_toSubsemiring {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {ι : Sort u_1} [Nonempty ι] (S : ιSubalgebra R A) :
          (iSup S).toSubsemiring = ⨆ (i : ι), (S i).toSubsemiring
          theorem Algebra.mem_iSup_of_mem {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {ι : Sort u_1} {S : ιSubalgebra R A} (i : ι) {x : A} (hx : x S i) :
          x iSup S
          theorem Algebra.iSup_induction {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {ι : Sort u_1} (S : ιSubalgebra R A) {motive : AProp} {x : A} (mem : x ⨆ (i : ι), S i) (basic : ∀ (i : ι), aS i, motive a) (zero : motive 0) (one : motive 1) (add : ∀ (a b : A), motive amotive bmotive (a + b)) (mul : ∀ (a b : A), motive amotive bmotive (a * b)) (algebraMap : ∀ (r : R), motive ((algebraMap R A) r)) :
          motive x
          theorem Algebra.iSup_induction' {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {ι : Sort u_1} (S : ιSubalgebra R A) {motive : (x : A) → x ⨆ (i : ι), S iProp} {x : A} (mem : x ⨆ (i : ι), S i) (basic : ∀ (i : ι) (x : A) (hx : x S i), motive x ) (zero : motive 0 ) (one : motive 1 ) (add : ∀ (x y : A) (hx : x ⨆ (i : ι), S i) (hy : y ⨆ (i : ι), S i), motive x hxmotive y hymotive (x + y) ) (mul : ∀ (x y : A) (hx : x ⨆ (i : ι), S i) (hy : y ⨆ (i : ι), S i), motive x hxmotive y hymotive (x * y) ) (algebraMap : ∀ (r : R), motive ((algebraMap R A) r) ) :
          motive x mem

          A dependent version of Subalgebra.iSup_induction.

          Equations
            theorem Algebra.mem_bot {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {x : A} :
            @[simp]
            theorem Algebra.coe_bot {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] :
            theorem Algebra.eq_top_iff {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {S : Subalgebra R A} :
            S = ∀ (x : A), x S
            theorem AlgHom.range_eq_top {R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A →ₐ[R] B) :
            @[deprecated AlgHom.range_eq_top (since := "2024-11-11")]
            theorem Algebra.range_top_iff_surjective {R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A →ₐ[R] B) :

            Alias of AlgHom.range_eq_top.

            @[simp]
            theorem Algebra.range_ofId {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] :
            (ofId R A).range =
            @[simp]
            theorem Algebra.range_id {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] :
            @[simp]
            theorem Algebra.map_top {R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A →ₐ[R] B) :
            @[simp]
            theorem Algebra.map_bot {R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A →ₐ[R] B) :
            @[simp]
            theorem Algebra.comap_top {R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A →ₐ[R] B) :
            def Algebra.toTop {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] :

            AlgHom to ⊤ : Subalgebra R A.

            Equations
              Instances For
                noncomputable def Algebra.botEquivOfInjective {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] (h : Function.Injective (algebraMap R A)) :

                The bottom subalgebra is isomorphic to the base ring.

                Equations
                  Instances For
                    noncomputable def Algebra.botEquiv (F : Type u_1) (R : Type u_2) [Field F] [Semiring R] [Nontrivial R] [Algebra F R] :

                    The bottom subalgebra is isomorphic to the field.

                    Equations
                      Instances For
                        @[simp]
                        theorem Algebra.botEquiv_symm_apply (F : Type u_1) (R : Type u_2) [Field F] [Semiring R] [Nontrivial R] [Algebra F R] (a✝ : F) :
                        (botEquiv F R).symm a✝ = (ofId F ) a✝
                        def Subalgebra.topEquiv {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] :

                        The top subalgebra is isomorphic to the algebra.

                        This is the algebra version of Submodule.topEquiv.

                        Equations
                          Instances For
                            @[simp]
                            theorem Subalgebra.topEquiv_symm_apply_coe {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] (a : A) :
                            (topEquiv.symm a) = a
                            @[simp]
                            theorem Subalgebra.topEquiv_apply {R : Type u} {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] (a : ) :
                            topEquiv a = a
                            instance AlgHom.subsingleton {R : Type u} {A : Type v} {B : Type w} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] [Subsingleton (Subalgebra R A)] :
                            Equations
                              @[simp]
                              theorem Subalgebra.center_eq_top (R : Type u) [CommSemiring R] (A : Type u_1) [CommSemiring A] [Algebra R A] :
                              @[simp]
                              theorem Subalgebra.centralizer_eq_top_iff_subset (R : Type u) {A : Type v} [CommSemiring R] [Semiring A] [Algebra R A] {s : Set A} :
                              centralizer R s = s (center R A)
                              @[simp]
                              theorem AlgHom.equalizer_eq_top {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {F : Type u_4} [FunLike F A B] [AlgHomClass F R A B] {φ ψ : F} :
                              equalizer φ ψ = φ = ψ
                              @[simp]
                              theorem AlgHom.equalizer_same {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {F : Type u_4} [FunLike F A B] [AlgHomClass F R A B] (φ : F) :
                              theorem AlgHom.eqOn_sup {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {F : Type u_4} [FunLike F A B] [AlgHomClass F R A B] {φ ψ : F} {S T : Subalgebra R A} (hS : Set.EqOn φ ψ S) (hT : Set.EqOn φ ψ T) :
                              Set.EqOn φ ψ (ST)
                              theorem AlgHom.ext_on_codisjoint {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {F : Type u_4} [FunLike F A B] [AlgHomClass F R A B] {φ ψ : F} {S T : Subalgebra R A} (hST : Codisjoint S T) (hS : Set.EqOn φ ψ S) (hT : Set.EqOn φ ψ T) :
                              φ = ψ
                              theorem Subalgebra.map_comap_eq {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] (f : A →ₐ[R] B) (S : Subalgebra R B) :
                              map f (comap f S) = Sf.range
                              theorem Subalgebra.map_comap_eq_self {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {f : A →ₐ[R] B} {S : Subalgebra R B} (h : S f.range) :
                              map f (comap f S) = S
                              theorem Subalgebra.map_comap_eq_self_of_surjective {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] [Algebra R B] {f : A →ₐ[R] B} (hf : Function.Surjective f) (S : Subalgebra R B) :
                              map f (comap f S) = S
                              theorem Algebra.subset_adjoin {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {s : Set A} :
                              s (adjoin R s)
                              theorem Algebra.adjoin_le {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {s : Set A} {S : Subalgebra R A} (H : s S) :
                              adjoin R s S
                              theorem Algebra.adjoin_singleton_le {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {S : Subalgebra R A} {a : A} (H : a S) :
                              theorem Algebra.adjoin_eq_sInf {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {s : Set A} :
                              adjoin R s = sInf {p : Subalgebra R A | s p}
                              theorem Algebra.adjoin_le_iff {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {s : Set A} {S : Subalgebra R A} :
                              adjoin R s S s S
                              theorem Algebra.adjoin_mono {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {s t : Set A} (H : s t) :
                              adjoin R s adjoin R t
                              theorem Algebra.adjoin_eq_of_le {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {s : Set A} (S : Subalgebra R A) (h₁ : s S) (h₂ : S adjoin R s) :
                              adjoin R s = S
                              theorem Algebra.adjoin_eq {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] (S : Subalgebra R A) :
                              adjoin R S = S
                              theorem Algebra.adjoin_iUnion {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {α : Type u_1} (s : αSet A) :
                              adjoin R (Set.iUnion s) = ⨆ (i : α), adjoin R (s i)
                              theorem Algebra.adjoin_attach_biUnion {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] [DecidableEq A] {α : Type u_1} {s : Finset α} (f : { x : α // x s }Finset A) :
                              adjoin R (s.attach.biUnion f) = ⨆ (x : { x : α // x s }), adjoin R (f x)
                              theorem Algebra.adjoin_induction {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {s : Set A} {p : (x : A) → x adjoin R sProp} (mem : ∀ (x : A) (hx : x s), p x ) (algebraMap : ∀ (r : R), p ((algebraMap R A) r) ) (add : ∀ (x y : A) (hx : x adjoin R s) (hy : y adjoin R s), p x hxp y hyp (x + y) ) (mul : ∀ (x y : A) (hx : x adjoin R s) (hy : y adjoin R s), p x hxp y hyp (x * y) ) {x : A} (hx : x adjoin R s) :
                              p x hx
                              theorem Algebra.adjoin_induction₂ {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {s : Set A} {p : (x y : A) → x adjoin R sy adjoin R sProp} (mem_mem : ∀ (x y : A) (hx : x s) (hy : y s), p x y ) (algebraMap_both : ∀ (r₁ r₂ : R), p ((algebraMap R A) r₁) ((algebraMap R A) r₂) ) (algebraMap_left : ∀ (r : R) (x : A) (hx : x s), p ((algebraMap R A) r) x ) (algebraMap_right : ∀ (r : R) (x : A) (hx : x s), p x ((algebraMap R A) r) ) (add_left : ∀ (x y z : A) (hx : x adjoin R s) (hy : y adjoin R s) (hz : z adjoin R s), p x z hx hzp y z hy hzp (x + y) z hz) (add_right : ∀ (x y z : A) (hx : x adjoin R s) (hy : y adjoin R s) (hz : z adjoin R s), p x y hx hyp x z hx hzp x (y + z) hx ) (mul_left : ∀ (x y z : A) (hx : x adjoin R s) (hy : y adjoin R s) (hz : z adjoin R s), p x z hx hzp y z hy hzp (x * y) z hz) (mul_right : ∀ (x y z : A) (hx : x adjoin R s) (hy : y adjoin R s) (hz : z adjoin R s), p x y hx hyp x z hx hzp x (y * z) hx ) {x y : A} (hx : x adjoin R s) (hy : y adjoin R s) :
                              p x y hx hy

                              Induction principle for the algebra generated by a set s: show that p x y holds for any x y ∈ adjoin R s given that it holds for x y ∈ s and that it satisfies a number of natural properties.

                              @[simp]
                              theorem Algebra.adjoin_adjoin_coe_preimage {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {s : Set A} :
                              theorem Algebra.adjoin_union {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] (s t : Set A) :
                              adjoin R (s t) = adjoin R sadjoin R t
                              @[simp]
                              theorem Algebra.adjoin_empty (R : Type uR) (A : Type uA) [CommSemiring R] [Semiring A] [Algebra R A] :
                              @[simp]
                              theorem Algebra.adjoin_univ (R : Type uR) (A : Type uA) [CommSemiring R] [Semiring A] [Algebra R A] :
                              @[simp]
                              theorem Algebra.adjoin_singleton_algebraMap {R : Type uR} (A : Type uA) [CommSemiring R] [Semiring A] [Algebra R A] (x : R) :
                              @[simp]
                              theorem Algebra.adjoin_singleton_natCast (R : Type uR) (A : Type uA) [CommSemiring R] [Semiring A] [Algebra R A] (n : ) :
                              adjoin R {n} =
                              @[simp]
                              theorem Algebra.adjoin_singleton_zero (R : Type uR) (A : Type uA) [CommSemiring R] [Semiring A] [Algebra R A] :
                              @[simp]
                              theorem Algebra.adjoin_singleton_one (R : Type uR) (A : Type uA) [CommSemiring R] [Semiring A] [Algebra R A] :
                              @[simp]
                              theorem Algebra.adjoin_insert_algebraMap {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] (x : R) (s : Set A) :
                              adjoin R (insert ((algebraMap R A) x) s) = adjoin R s
                              @[simp]
                              theorem Algebra.adjoin_insert_natCast (R : Type uR) {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] (n : ) (s : Set A) :
                              adjoin R (insert (↑n) s) = adjoin R s
                              @[simp]
                              theorem Algebra.adjoin_insert_zero (R : Type uR) {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] (s : Set A) :
                              adjoin R (insert 0 s) = adjoin R s
                              @[simp]
                              theorem Algebra.adjoin_insert_one (R : Type uR) {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] (s : Set A) :
                              adjoin R (insert 1 s) = adjoin R s
                              theorem Algebra.adjoin_toSubmodule_le (R : Type uR) {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {s : Set A} {t : Submodule R A} :
                              @[simp]
                              theorem Algebra.adjoin_span (R : Type uR) {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {s : Set A} :
                              adjoin R (Submodule.span R s) = adjoin R s
                              theorem Algebra.adjoin_image (R : Type uR) {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (f : A →ₐ[R] B) (s : Set A) :
                              adjoin R (f '' s) = Subalgebra.map f (adjoin R s)
                              @[simp]
                              theorem Algebra.adjoin_insert_adjoin (R : Type uR) {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] (s : Set A) (x : A) :
                              adjoin R (insert x (adjoin R s)) = adjoin R (insert x s)
                              theorem Algebra.mem_adjoin_of_map_mul (R : Type uR) {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] {s : Set A} {x : A} {f : A →ₗ[R] B} (hf : ∀ (a₁ a₂ : A), f (a₁ * a₂) = f a₁ * f a₂) (h : x adjoin R s) :
                              f x adjoin R (f '' (s {1}))
                              @[reducible, inline]
                              abbrev Algebra.adjoinCommSemiringOfComm (R : Type uR) {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {s : Set A} (hcomm : as, bs, a * b = b * a) :

                              If all elements of s : Set A commute pairwise, then adjoin s is a commutative semiring.

                              Equations
                                Instances For
                                  theorem Algebra.commute_of_mem_adjoin_of_forall_mem_commute {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {a b : A} {s : Set A} (hb : b adjoin R s) (h : bs, Commute a b) :
                                  theorem Algebra.commute_of_mem_adjoin_singleton_of_commute {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {a b c : A} (hc : c adjoin R {b}) (h : Commute a b) :
                                  theorem Algebra.commute_of_mem_adjoin_self {R : Type uR} {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] {a b : A} (hb : b adjoin R {a}) :
                                  theorem Algebra.self_mem_adjoin_singleton (R : Type uR) {A : Type uA} [CommSemiring R] [Semiring A] [Algebra R A] (x : A) :
                                  @[simp]
                                  theorem Algebra.adjoin_singleton_intCast {R : Type uR} {A : Type uA} [CommRing R] [Ring A] [Algebra R A] (n : ) :
                                  adjoin R {n} =
                                  @[simp]
                                  theorem Algebra.adjoin_insert_intCast {R : Type uR} {A : Type uA} [CommRing R] [Ring A] [Algebra R A] (n : ) (s : Set A) :
                                  adjoin R (insert (↑n) s) = adjoin R s
                                  theorem Algebra.mem_adjoin_iff {R : Type uR} {A : Type uA} [CommRing R] [Ring A] [Algebra R A] {s : Set A} {x : A} :
                                  theorem Algebra.adjoin_eq_ring_closure {R : Type uR} {A : Type uA} [CommRing R] [Ring A] [Algebra R A] (s : Set A) :
                                  @[reducible, inline]
                                  abbrev Algebra.adjoinCommRingOfComm (R : Type uR) {A : Type uA} [CommRing R] [Ring A] [Algebra R A] {s : Set A} (hcomm : as, bs, a * b = b * a) :
                                  CommRing (adjoin R s)

                                  If all elements of s : Set A commute pairwise, then adjoin R s is a commutative ring.

                                  Equations
                                    Instances For
                                      theorem AlgHom.map_adjoin {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (φ : A →ₐ[R] B) (s : Set A) :
                                      @[simp]
                                      theorem AlgHom.map_adjoin_singleton {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (e : A →ₐ[R] B) (x : A) :
                                      theorem AlgHom.adjoin_le_equalizer {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (φ₁ φ₂ : A →ₐ[R] B) {s : Set A} (h : Set.EqOn (⇑φ₁) (⇑φ₂) s) :
                                      Algebra.adjoin R s equalizer φ₁ φ₂
                                      theorem AlgHom.ext_of_adjoin_eq_top {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] {s : Set A} (h : Algebra.adjoin R s = ) φ₁ φ₂ : A →ₐ[R] B (hs : Set.EqOn (⇑φ₁) (⇑φ₂) s) :
                                      φ₁ = φ₂
                                      theorem AlgHom.eqOn_adjoin_iff {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] {φ ψ : A →ₐ[R] B} {s : Set A} :
                                      Set.EqOn φ ψ (Algebra.adjoin R s) Set.EqOn (⇑φ) (⇑ψ) s

                                      Two algebra morphisms are equal on Algebra.span siff they are equal on s

                                      theorem AlgHom.adjoin_ext {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] {s : Set A} φ₁ φ₂ : (Algebra.adjoin R s) →ₐ[R] B (h : ∀ (x : A) (hx : x s), φ₁ x, = φ₂ x, ) :
                                      φ₁ = φ₂
                                      theorem AlgHom.ext_of_eq_adjoin {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] {S : Subalgebra R A} {s : Set A} (hS : S = Algebra.adjoin R s) φ₁ φ₂ : S →ₐ[R] B (h : ∀ (x : A) (hx : x s), φ₁ x, = φ₂ x, ) :
                                      φ₁ = φ₂

                                      The -algebra equivalence between Subsemiring.closure s and Algebra.adjoin ℕ s given by the identity map.

                                      Equations
                                        Instances For

                                          The -algebra equivalence between Subring.closure s and Algebra.adjoin ℤ s given by the identity map.

                                          Equations
                                            Instances For
                                              theorem Submonoid.adjoin_eq_span_of_eq_span (F : Type u_1) (E : Type u_2) {K : Type u_3} [CommSemiring E] [Semiring K] [SMul F E] [Algebra E K] [Semiring F] [Module F K] [IsScalarTower F E K] (L : Submonoid K) {S : Set K} (h : L = (Submodule.span F S)) :

                                              If K / E / F is a ring extension tower, L is a submonoid of K / F which is generated by S as an F-module, then E[L] is generated by S as an E-module.

                                              theorem Subalgebra.adjoin_eq_span_of_eq_span {F : Type u_1} (E : Type u_2) {K : Type u_3} [CommSemiring E] [Semiring K] [SMul F E] [Algebra E K] [CommSemiring F] [Algebra F K] [IsScalarTower F E K] (L : Subalgebra F K) {S : Set K} (h : toSubmodule L = Submodule.span F S) :

                                              If K / E / F is a ring extension tower, L is a subalgebra of K / F which is generated by S as an F-module, then E[L] is generated by S as an E-module.

                                              theorem Algebra.adjoin_nonUnitalSubalgebra (R : Type uR) {A : Type uA} [CommSemiring R] [Ring A] [Algebra R A] (s : Set A) :
                                              theorem Subalgebra.comap_map_eq {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Ring A] [Algebra R A] [Ring B] [Algebra R B] (f : A →ₐ[R] B) (S : Subalgebra R A) :
                                              comap f (map f S) = SAlgebra.adjoin R (f ⁻¹' {0})
                                              theorem Subalgebra.comap_map_eq_self {R : Type uR} {A : Type uA} {B : Type uB} [CommSemiring R] [Ring A] [Algebra R A] [Ring B] [Algebra R B] {f : A →ₐ[R] B} {S : Subalgebra R A} (h : f ⁻¹' {0} S) :
                                              comap f (map f S) = S