Documentation

Mathlib.RingTheory.Congruence.Defs

Congruence relations on rings #

This file defines congruence relations on rings, which extend Con and AddCon on monoids and additive monoids.

Most of the time you likely want to use the Ideal.Quotient API that is built on top of this.

Main Definitions #

TODO #

structure RingCon (R : Type u_1) [Add R] [Mul R] extends Con R, AddCon R :
Type u_1

A congruence relation on a type with an addition and multiplication is an equivalence relation which preserves both.

  • r : RRProp
  • iseqv : Equivalence self.toSetoid
  • mul' : ∀ {w x y z : R}, self.toSetoid w xself.toSetoid y zself.toSetoid (w * y) (x * z)
  • add' : ∀ {w x y z : R}, self.toSetoid w xself.toSetoid y zself.toSetoid (w + y) (x + z)
Instances For
    inductive RingConGen.Rel {R : Type u_1} [Add R] [Mul R] (r : RRProp) :
    RRProp

    The inductively defined smallest ring congruence relation containing a given binary relation.

    Instances For
      def ringConGen {R : Type u_1} [Add R] [Mul R] (r : RRProp) :

      The inductively defined smallest ring congruence relation containing a given binary relation.

      Equations
      Instances For
        instance RingCon.instFunLikeForallProp {R : Type u_1} [Add R] [Mul R] :
        FunLike (RingCon R) R (RProp)

        A coercion from a congruence relation to its underlying binary relation.

        Equations
        • RingCon.instFunLikeForallProp = { coe := fun (c : RingCon R) => c.toSetoid, coe_injective' := }
        theorem RingCon.rel_eq_coe {R : Type u_1} [Add R] [Mul R] (c : RingCon R) :
        c.toSetoid = c
        @[simp]
        theorem RingCon.toCon_coe_eq_coe {R : Type u_1} [Add R] [Mul R] (c : RingCon R) :
        c.toCon = c
        theorem RingCon.refl {R : Type u_1} [Add R] [Mul R] (c : RingCon R) (x : R) :
        c x x
        theorem RingCon.symm {R : Type u_1} [Add R] [Mul R] (c : RingCon R) {x y : R} :
        c x yc y x
        theorem RingCon.trans {R : Type u_1} [Add R] [Mul R] (c : RingCon R) {x y z : R} :
        c x yc y zc x z
        theorem RingCon.add {R : Type u_1} [Add R] [Mul R] (c : RingCon R) {w x y z : R} :
        c w xc y zc (w + y) (x + z)
        theorem RingCon.mul {R : Type u_1} [Add R] [Mul R] (c : RingCon R) {w x y z : R} :
        c w xc y zc (w * y) (x * z)
        theorem RingCon.sub {S : Type u_2} [AddGroup S] [Mul S] (t : RingCon S) {a b c d : S} (h : t a b) (h' : t c d) :
        t (a - c) (b - d)
        theorem RingCon.neg {S : Type u_2} [AddGroup S] [Mul S] (t : RingCon S) {a b : S} (h : t a b) :
        t (-a) (-b)
        theorem RingCon.nsmul {S : Type u_2} [AddGroup S] [Mul S] (t : RingCon S) (m : ) {x y : S} (hx : t x y) :
        t (m x) (m y)
        theorem RingCon.zsmul {S : Type u_2} [AddGroup S] [Mul S] (t : RingCon S) (z : ) {x y : S} (hx : t x y) :
        t (z x) (z y)
        instance RingCon.instInhabited {R : Type u_1} [Add R] [Mul R] :
        Equations
        • RingCon.instInhabited = { default := ringConGen EmptyRelation }
        @[simp]
        theorem RingCon.rel_mk {R : Type u_1} [Add R] [Mul R] {s : Con R} {h : ∀ {w x y z : R}, s.toSetoid w xs.toSetoid y zs.toSetoid (w + y) (x + z)} {a b : R} :
        { toCon := s, add' := h } a b s a b
        theorem RingCon.ext' {R : Type u_1} [Add R] [Mul R] {c d : RingCon R} (H : c = d) :
        c = d

        The map sending a congruence relation to its underlying binary relation is injective.

        theorem RingCon.ext {R : Type u_1} [Add R] [Mul R] {c d : RingCon R} (H : ∀ (x y : R), c x y d x y) :
        c = d

        Extensionality rule for congruence relations.

        def RingCon.comap {R : Type u_2} {R' : Type u_3} {F : Type u_4} [Add R] [Add R'] [FunLike F R R'] [AddHomClass F R R'] [Mul R] [Mul R'] [MulHomClass F R R'] (J : RingCon R') (f : F) :

        Pulling back a RingCon across a ring homomorphism.

        Equations
        • J.comap f = { toCon := Con.comap f J.toCon, add' := }
        Instances For
          def RingCon.Quotient {R : Type u_1} [Add R] [Mul R] (c : RingCon R) :
          Type u_1

          Defining the quotient by a congruence relation of a type with addition and multiplication.

          Equations
          Instances For
            def RingCon.toQuotient {R : Type u_1} [Add R] [Mul R] {c : RingCon R} (r : R) :
            c.Quotient

            The morphism into the quotient by a congruence relation

            Equations
            Instances For
              instance RingCon.instCoeTCQuotient {R : Type u_1} [Add R] [Mul R] (c : RingCon R) :
              CoeTC R c.Quotient

              Coercion from a type with addition and multiplication to its quotient by a congruence relation.

              See Note [use has_coe_t].

              Equations
              • c.instCoeTCQuotient = { coe := RingCon.toQuotient }
              @[instance 500]
              instance RingCon.instDecidableEqQuotientOfDecidableCoeForallProp {R : Type u_1} [Add R] [Mul R] (c : RingCon R) [_d : (a b : R) → Decidable (c a b)] :
              DecidableEq c.Quotient

              The quotient by a decidable congruence relation has decidable equality.

              Equations
              @[simp]
              theorem RingCon.quot_mk_eq_coe {R : Type u_1} [Add R] [Mul R] (c : RingCon R) (x : R) :
              Quot.mk (⇑c) x = x
              @[simp]
              theorem RingCon.eq {R : Type u_1} [Add R] [Mul R] (c : RingCon R) {a b : R} :
              a = b c a b

              Two elements are related by a congruence relation c iff they are represented by the same element of the quotient by c.

              Basic notation #

              The basic algebraic notation, 0, 1, +, *, -, ^, descend naturally under the quotient

              instance RingCon.instAddQuotient {R : Type u_1} [Add R] [Mul R] (c : RingCon R) :
              Add c.Quotient
              Equations
              @[simp]
              theorem RingCon.coe_add {R : Type u_1} [Add R] [Mul R] (c : RingCon R) (x y : R) :
              (x + y) = x + y
              instance RingCon.instMulQuotient {R : Type u_1} [Add R] [Mul R] (c : RingCon R) :
              Mul c.Quotient
              Equations
              @[simp]
              theorem RingCon.coe_mul {R : Type u_1} [Add R] [Mul R] (c : RingCon R) (x y : R) :
              (x * y) = x * y
              instance RingCon.instZeroQuotient {R : Type u_1} [AddZeroClass R] [Mul R] (c : RingCon R) :
              Zero c.Quotient
              Equations
              @[simp]
              theorem RingCon.coe_zero {R : Type u_1} [AddZeroClass R] [Mul R] (c : RingCon R) :
              0 = 0
              instance RingCon.instOneQuotient {R : Type u_1} [Add R] [MulOneClass R] (c : RingCon R) :
              One c.Quotient
              Equations
              @[simp]
              theorem RingCon.coe_one {R : Type u_1} [Add R] [MulOneClass R] (c : RingCon R) :
              1 = 1
              instance RingCon.instNegQuotient {R : Type u_1} [AddGroup R] [Mul R] (c : RingCon R) :
              Neg c.Quotient
              Equations
              @[simp]
              theorem RingCon.coe_neg {R : Type u_1} [AddGroup R] [Mul R] (c : RingCon R) (x : R) :
              (-x) = -x
              instance RingCon.instSubQuotient {R : Type u_1} [AddGroup R] [Mul R] (c : RingCon R) :
              Sub c.Quotient
              Equations
              @[simp]
              theorem RingCon.coe_sub {R : Type u_1} [AddGroup R] [Mul R] (c : RingCon R) (x y : R) :
              (x - y) = x - y
              instance RingCon.hasZSMul {R : Type u_1} [AddGroup R] [Mul R] (c : RingCon R) :
              SMul c.Quotient
              Equations
              @[simp]
              theorem RingCon.coe_zsmul {R : Type u_1} [AddGroup R] [Mul R] (c : RingCon R) (z : ) (x : R) :
              (z x) = z x
              instance RingCon.hasNSMul {R : Type u_1} [AddMonoid R] [Mul R] (c : RingCon R) :
              SMul c.Quotient
              Equations
              @[simp]
              theorem RingCon.coe_nsmul {R : Type u_1} [AddMonoid R] [Mul R] (c : RingCon R) (n : ) (x : R) :
              (n x) = n x
              instance RingCon.instPowQuotientNat {R : Type u_1} [Add R] [Monoid R] (c : RingCon R) :
              Pow c.Quotient
              Equations
              @[simp]
              theorem RingCon.coe_pow {R : Type u_1} [Add R] [Monoid R] (c : RingCon R) (x : R) (n : ) :
              (x ^ n) = x ^ n
              instance RingCon.instNatCastQuotient {R : Type u_1} [AddMonoidWithOne R] [Mul R] (c : RingCon R) :
              NatCast c.Quotient
              Equations
              • c.instNatCastQuotient = { natCast := fun (n : ) => n }
              @[simp]
              theorem RingCon.coe_natCast {R : Type u_1} [AddMonoidWithOne R] [Mul R] (c : RingCon R) (n : ) :
              n = n
              @[deprecated RingCon.coe_natCast]
              theorem RingCon.coe_nat_cast {R : Type u_1} [AddMonoidWithOne R] [Mul R] (c : RingCon R) (n : ) :
              n = n

              Alias of RingCon.coe_natCast.

              instance RingCon.instIntCastQuotient {R : Type u_1} [AddGroupWithOne R] [Mul R] (c : RingCon R) :
              IntCast c.Quotient
              Equations
              • c.instIntCastQuotient = { intCast := fun (z : ) => z }
              @[simp]
              theorem RingCon.coe_intCast {R : Type u_1} [AddGroupWithOne R] [Mul R] (c : RingCon R) (n : ) :
              n = n
              @[deprecated RingCon.coe_intCast]
              theorem RingCon.coe_int_cast {R : Type u_1} [AddGroupWithOne R] [Mul R] (c : RingCon R) (n : ) :
              n = n

              Alias of RingCon.coe_intCast.

              instance RingCon.instInhabitedQuotient {R : Type u_1} [Inhabited R] [Add R] [Mul R] (c : RingCon R) :
              Inhabited c.Quotient
              Equations
              • c.instInhabitedQuotient = { default := default }

              Algebraic structure #

              The operations above on the quotient by c : RingCon R preserve the algebraic structure of R.

              Equations
              Equations
              Equations
              instance RingCon.instSemiringQuotient {R : Type u_1} [Semiring R] (c : RingCon R) :
              Semiring c.Quotient
              Equations
              instance RingCon.instCommSemiringQuotient {R : Type u_1} [CommSemiring R] (c : RingCon R) :
              CommSemiring c.Quotient
              Equations
              Equations
              instance RingCon.instNonAssocRingQuotient {R : Type u_1} [NonAssocRing R] (c : RingCon R) :
              NonAssocRing c.Quotient
              Equations
              instance RingCon.instNonUnitalRingQuotient {R : Type u_1} [NonUnitalRing R] (c : RingCon R) :
              NonUnitalRing c.Quotient
              Equations
              instance RingCon.instRingQuotient {R : Type u_1} [Ring R] (c : RingCon R) :
              Ring c.Quotient
              Equations
              instance RingCon.instCommRingQuotient {R : Type u_1} [CommRing R] (c : RingCon R) :
              CommRing c.Quotient
              Equations
              def RingCon.mk' {R : Type u_1} [NonAssocSemiring R] (c : RingCon R) :
              R →+* c.Quotient

              The natural homomorphism from a ring to its quotient by a congruence relation.

              Equations
              • c.mk' = { toFun := RingCon.toQuotient, map_one' := , map_mul' := , map_zero' := , map_add' := }
              Instances For