Documentation

Mathlib.Algebra.Ring.Subring.Order

Subrings of ordered rings #

We study subrings of ordered rings and prove their basic properties.

Main definitions and results #

@[instance 75]
instance SubringClass.toOrderedRing {R : Type u_1} {S : Type u_2} [SetLike S R] (s : S) [OrderedRing R] [SubringClass S R] :

A subring of an OrderedRing is an OrderedRing.

Equations
@[instance 75]
instance SubringClass.toOrderedCommRing {R : Type u_1} {S : Type u_2} [SetLike S R] (s : S) [OrderedCommRing R] [SubringClass S R] :

A subring of an OrderedCommRing is an OrderedCommRing.

Equations
@[instance 75]
instance SubringClass.toLinearOrderedRing {R : Type u_1} {S : Type u_2} [SetLike S R] (s : S) [LinearOrderedRing R] [SubringClass S R] :

A subring of a LinearOrderedRing is a LinearOrderedRing.

Equations
@[instance 75]

A subring of a LinearOrderedCommRing is a LinearOrderedCommRing.

Equations
instance Subring.toOrderedRing {R : Type u_1} [OrderedRing R] (s : Subring R) :

A subring of an OrderedRing is an OrderedRing.

Equations

A subring of an OrderedCommRing is an OrderedCommRing.

Equations

A subring of a LinearOrderedRing is a LinearOrderedRing.

Equations
def Subring.orderedSubtype {R : Type u_2} [OrderedRing R] (s : Subring R) :
s →+*o R

The inclusion S → R of a subring, as an ordered ring homomorphism.

Equations
  • s.orderedSubtype = let __spread.0 := s.subtype; { toRingHom := __spread.0, monotone' := }
Instances For
    theorem Subring.orderedSubtype_coe {R : Type u_2} [OrderedRing R] (s : Subring R) :
    s.orderedSubtype = s.subtype