Documentation

Mathlib.GroupTheory.Congruence.Basic

Congruence relations #

This file defines congruence relations: equivalence relations that preserve a binary operation, which in this case is multiplication or addition. The principal definition is a structure extending a Setoid (an equivalence relation), and the inductive definition of the smallest congruence relation containing a binary relation is also given (see ConGen).

The file also proves basic properties of the quotient of a type by a congruence relation, and the complete lattice of congruence relations on a type. We then establish an order-preserving bijection between the set of congruence relations containing a congruence relation c and the set of congruence relations on the quotient by c.

The second half of the file concerns congruence relations on monoids, in which case the quotient by the congruence relation is also a monoid. There are results about the universal property of quotients of monoids, and the isomorphism theorems for monoids.

Implementation notes #

The inductive definition of a congruence relation could be a nested inductive type, defined using the equivalence closure of a binary relation EqvGen, but the recursor generated does not work. A nested inductive definition could conceivably shorten proofs, because they would allow invocation of the corresponding lemmas about EqvGen.

The lemmas refl, symm and trans are not tagged with @[refl], @[symm], and @[trans] respectively as these tags do not work on a structure coerced to a binary relation.

There is a coercion from elements of a type to the element's equivalence class under a congruence relation.

A congruence relation on a monoid M can be thought of as a submonoid of M × M for which membership is an equivalence relation, but whilst this fact is established in the file, it is not used, since this perspective adds more layers of definitional unfolding.

Tags #

congruence, congruence relation, quotient, quotient by congruence relation, monoid, quotient monoid, isomorphism theorems

structure AddCon (M : Type u_1) [Add M] extends Setoid :
Type u_1

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

Instances For
theorem AddCon.add' {M : Type u_1} [Add M] (self : AddCon M) {w : M} {x : M} {y : M} {z : M} :
Setoid.r w xSetoid.r y zSetoid.r (w + y) (x + z)

Additive congruence relations are closed under addition

structure Con (M : Type u_1) [Mul M] extends Setoid :
Type u_1

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

Instances For
theorem Con.mul' {M : Type u_1} [Mul M] (self : Con M) {w : M} {x : M} {y : M} {z : M} :
Setoid.r w xSetoid.r y zSetoid.r (w * y) (x * z)

Congruence relations are closed under multiplication

inductive AddConGen.Rel {M : Type u_1} [Add M] (r : MMProp) :
MMProp

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

inductive ConGen.Rel {M : Type u_1} [Mul M] (r : MMProp) :
MMProp

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

theorem addConGen.proof_1 {M : Type u_1} [Add M] (r : MMProp) :
def addConGen {M : Type u_1} [Add M] (r : MMProp) :

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

Equations
def conGen {M : Type u_1} [Mul M] (r : MMProp) :
Con M

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

Equations
instance AddCon.instInhabited {M : Type u_1} [Add M] :
Equations
  • AddCon.instInhabited = { default := addConGen EmptyRelation }
instance Con.instInhabited {M : Type u_1} [Mul M] :
Equations
  • Con.instInhabited = { default := conGen EmptyRelation }
theorem AddCon.instFunLikeForallProp.proof_1 {M : Type u_1} [Add M] (x : AddCon M) (y : AddCon M) (h : (fun (c : AddCon M) => Setoid.r) x = (fun (c : AddCon M) => Setoid.r) y) :
x = y
instance AddCon.instFunLikeForallProp {M : Type u_1} [Add M] :
FunLike (AddCon M) M (MProp)

A coercion from an additive congruence relation to its underlying binary relation.

Equations
  • AddCon.instFunLikeForallProp = { coe := fun (c : AddCon M) => Setoid.r, coe_injective' := }
instance Con.instFunLikeForallProp {M : Type u_1} [Mul M] :
FunLike (Con M) M (MProp)

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

Equations
  • Con.instFunLikeForallProp = { coe := fun (c : Con M) => Setoid.r, coe_injective' := }
@[simp]
theorem AddCon.rel_eq_coe {M : Type u_1} [Add M] (c : AddCon M) :
Setoid.r = c
@[simp]
theorem Con.rel_eq_coe {M : Type u_1} [Mul M] (c : Con M) :
Setoid.r = c
theorem AddCon.refl {M : Type u_1} [Add M] (c : AddCon M) (x : M) :
c x x

Additive congruence relations are reflexive.

theorem Con.refl {M : Type u_1} [Mul M] (c : Con M) (x : M) :
c x x

Congruence relations are reflexive.

theorem AddCon.symm {M : Type u_1} [Add M] (c : AddCon M) {x : M} {y : M} :
c x yc y x

Additive congruence relations are symmetric.

theorem Con.symm {M : Type u_1} [Mul M] (c : Con M) {x : M} {y : M} :
c x yc y x

Congruence relations are symmetric.

theorem AddCon.trans {M : Type u_1} [Add M] (c : AddCon M) {x : M} {y : M} {z : M} :
c x yc y zc x z

Additive congruence relations are transitive.

theorem Con.trans {M : Type u_1} [Mul M] (c : Con M) {x : M} {y : M} {z : M} :
c x yc y zc x z

Congruence relations are transitive.

theorem AddCon.add {M : Type u_1} [Add M] (c : AddCon M) {w : M} {x : M} {y : M} {z : M} :
c w xc y zc (w + y) (x + z)

Additive congruence relations preserve addition.

theorem Con.mul {M : Type u_1} [Mul M] (c : Con M) {w : M} {x : M} {y : M} {z : M} :
c w xc y zc (w * y) (x * z)

Multiplicative congruence relations preserve multiplication.

@[simp]
theorem AddCon.rel_mk {M : Type u_1} [Add M] {s : Setoid M} {h : ∀ {w x y z : M}, Setoid.r w xSetoid.r y zSetoid.r (w + y) (x + z)} {a : M} {b : M} :
{ toSetoid := s, add' := h } a b Setoid.r a b
@[simp]
theorem Con.rel_mk {M : Type u_1} [Mul M] {s : Setoid M} {h : ∀ {w x y z : M}, Setoid.r w xSetoid.r y zSetoid.r (w * y) (x * z)} {a : M} {b : M} :
{ toSetoid := s, mul' := h } a b Setoid.r a b
instance AddCon.instMembershipSum {M : Type u_1} [Add M] :
Membership (M × M) (AddCon M)

Given a type M with an addition, x, y ∈ M, and an additive congruence relation c on M, (x, y) ∈ M × M iff x is related to y by c.

Equations
  • AddCon.instMembershipSum = { mem := fun (c : AddCon M) (x : M × M) => c x.1 x.2 }
instance Con.instMembershipProd {M : Type u_1} [Mul M] :
Membership (M × M) (Con M)

Given a type M with a multiplication, a congruence relation c on M, and elements of M x, y, (x, y) ∈ M × M iff x is related to y by c.

Equations
  • Con.instMembershipProd = { mem := fun (c : Con M) (x : M × M) => c x.1 x.2 }
theorem AddCon.ext' {M : Type u_1} [Add M] {c : AddCon M} {d : AddCon M} (H : c = d) :
c = d

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

theorem Con.ext' {M : Type u_1} [Mul M] {c : Con M} {d : Con M} (H : c = d) :
c = d

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

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

Extensionality rule for additive congruence relations.

theorem Con.ext_iff {M : Type u_1} [Mul M] {c : Con M} {d : Con M} :
c = d ∀ (x y : M), c x y d x y
theorem AddCon.ext_iff {M : Type u_1} [Add M] {c : AddCon M} {d : AddCon M} :
c = d ∀ (x y : M), c x y d x y
theorem Con.ext {M : Type u_1} [Mul M] {c : Con M} {d : Con M} (H : ∀ (x y : M), c x y d x y) :
c = d

Extensionality rule for congruence relations.

theorem AddCon.toSetoid_inj {M : Type u_1} [Add M] {c : AddCon M} {d : AddCon M} (H : c.toSetoid = d.toSetoid) :
c = d

The map sending an additive congruence relation to its underlying equivalence relation is injective.

theorem Con.toSetoid_inj {M : Type u_1} [Mul M] {c : Con M} {d : Con M} (H : c.toSetoid = d.toSetoid) :
c = d

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

theorem AddCon.coe_inj {M : Type u_1} [Add M] {c : AddCon M} {d : AddCon M} :
c = d c = d

Two additive congruence relations are equal iff their underlying binary relations are equal.

theorem Con.coe_inj {M : Type u_1} [Mul M] {c : Con M} {d : Con M} :
c = d c = d

Two congruence relations are equal iff their underlying binary relations are equal.

def AddCon.addKer {M : Type u_1} {P : Type u_3} [Add M] [Add P] (f : MP) (h : ∀ (x y : M), f (x + y) = f x + f y) :

The kernel of an addition-preserving function as an additive congruence relation.

Equations
theorem AddCon.addKer.proof_1 {M : Type u_1} {P : Type u_2} [Add M] [Add P] (f : MP) (h : ∀ (x y : M), f (x + y) = f x + f y) :
∀ {w x y z : M}, Setoid.r w xSetoid.r y zSetoid.r (w + y) (x + z)
def Con.mulKer {M : Type u_1} {P : Type u_3} [Mul M] [Mul P] (f : MP) (h : ∀ (x y : M), f (x * y) = f x * f y) :
Con M

The kernel of a multiplication-preserving function as a congruence relation.

Equations
def AddCon.prod {M : Type u_1} {N : Type u_2} [Add M] [Add N] (c : AddCon M) (d : AddCon N) :
AddCon (M × N)

Given types with additions M, N, the product of two congruence relations c on M and d on N: (x₁, x₂), (y₁, y₂) ∈ M × N are related by c.prod d iff x₁ is related to y₁ by c and x₂ is related to y₂ by d.

Equations
  • c.prod d = { toSetoid := c.prod d.toSetoid, add' := }
theorem AddCon.prod.proof_1 {M : Type u_1} {N : Type u_2} [Add M] [Add N] (c : AddCon M) (d : AddCon N) :
∀ {w x y z : M × N}, Setoid.r w xSetoid.r y zc.Rel (w + y).1 (x + z).1 d.Rel (w + y).2 (x + z).2
def Con.prod {M : Type u_1} {N : Type u_2} [Mul M] [Mul N] (c : Con M) (d : Con N) :
Con (M × N)

Given types with multiplications M, N, the product of two congruence relations c on M and d on N: (x₁, x₂), (y₁, y₂) ∈ M × N are related by c.prod d iff x₁ is related to y₁ by c and x₂ is related to y₂ by d.

Equations
  • c.prod d = { toSetoid := c.prod d.toSetoid, mul' := }
theorem AddCon.pi.proof_1 {ι : Type u_1} {f : ιType u_2} [(i : ι) → Add (f i)] (C : (i : ι) → AddCon (f i)) :
∀ {w x y z : (i : ι) → f i}, Setoid.r w xSetoid.r y z∀ (i : ι), (C i) (w i + y i) (x i + z i)
def AddCon.pi {ι : Type u_4} {f : ιType u_5} [(i : ι) → Add (f i)] (C : (i : ι) → AddCon (f i)) :
AddCon ((i : ι) → f i)

The product of an indexed collection of additive congruence relations.

Equations
  • AddCon.pi C = { toSetoid := piSetoid, add' := }
def Con.pi {ι : Type u_4} {f : ιType u_5} [(i : ι) → Mul (f i)] (C : (i : ι) → Con (f i)) :
Con ((i : ι) → f i)

The product of an indexed collection of congruence relations.

Equations
  • Con.pi C = { toSetoid := piSetoid, mul' := }
def AddCon.toQuotient {M : Type u_1} [Add M] {c : AddCon M} :
Mc.Quotient

The morphism into the quotient by an additive congruence relation

Equations
  • AddCon.toQuotient = Quotient.mk''
def Con.toQuotient {M : Type u_1} [Mul M] {c : Con M} :
Mc.Quotient

The morphism into the quotient by a congruence relation

Equations
  • Con.toQuotient = Quotient.mk''
@[instance 10]
instance AddCon.instCoeTCQuotient {M : Type u_1} [Add M] (c : AddCon M) :
CoeTC M c.Quotient

Coercion from a type with an addition to its quotient by an additive congruence relation

Equations
  • c.instCoeTCQuotient = { coe := AddCon.toQuotient }
@[instance 10]
instance Con.instCoeTCQuotient {M : Type u_1} [Mul M] (c : Con M) :
CoeTC M c.Quotient

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

See Note [use has_coe_t].

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

The quotient by a decidable additive congruence relation has decidable equality.

Equations
@[instance 500]
instance Con.instDecidableEqQuotientOfDecidableCoeForallProp {M : Type u_1} [Mul M] (c : Con M) [(a b : M) → Decidable (c a b)] :
DecidableEq c.Quotient

The quotient by a decidable congruence relation has decidable equality.

Equations
@[simp]
theorem AddCon.quot_mk_eq_coe {M : Type u_4} [Add M] (c : AddCon M) (x : M) :
Quot.mk (⇑c) x = x
@[simp]
theorem Con.quot_mk_eq_coe {M : Type u_4} [Mul M] (c : Con M) (x : M) :
Quot.mk (⇑c) x = x
def AddCon.liftOn {M : Type u_1} [Add M] {β : Sort u_4} {c : AddCon M} (q : c.Quotient) (f : Mβ) (h : ∀ (a b : M), c a bf a = f b) :
β

The function on the quotient by a congruence relation c induced by a function that is constant on c's equivalence classes.

Equations
def Con.liftOn {M : Type u_1} [Mul M] {β : Sort u_4} {c : Con M} (q : c.Quotient) (f : Mβ) (h : ∀ (a b : M), c a bf a = f b) :
β

The function on the quotient by a congruence relation c induced by a function that is constant on c's equivalence classes.

Equations
def AddCon.liftOn₂ {M : Type u_1} [Add M] {β : Sort u_4} {c : AddCon M} (q : c.Quotient) (r : c.Quotient) (f : MMβ) (h : ∀ (a₁ a₂ b₁ b₂ : M), c a₁ b₁c a₂ b₂f a₁ a₂ = f b₁ b₂) :
β

The binary function on the quotient by a congruence relation c induced by a binary function that is constant on c's equivalence classes.

Equations
def Con.liftOn₂ {M : Type u_1} [Mul M] {β : Sort u_4} {c : Con M} (q : c.Quotient) (r : c.Quotient) (f : MMβ) (h : ∀ (a₁ a₂ b₁ b₂ : M), c a₁ b₁c a₂ b₂f a₁ a₂ = f b₁ b₂) :
β

The binary function on the quotient by a congruence relation c induced by a binary function that is constant on c's equivalence classes.

Equations
def AddCon.hrecOn₂ {M : Type u_1} {N : Type u_2} [Add M] [Add N] {cM : AddCon M} {cN : AddCon N} {φ : cM.QuotientcN.QuotientSort u_4} (a : cM.Quotient) (b : cN.Quotient) (f : (x : M) → (y : N) → φ x y) (h : ∀ (x : M) (y : N) (x' : M) (y' : N), cM x x'cN y y'HEq (f x y) (f x' y')) :
φ a b

A version of Quotient.hrecOn₂' for quotients by AddCon.

Equations
def Con.hrecOn₂ {M : Type u_1} {N : Type u_2} [Mul M] [Mul N] {cM : Con M} {cN : Con N} {φ : cM.QuotientcN.QuotientSort u_4} (a : cM.Quotient) (b : cN.Quotient) (f : (x : M) → (y : N) → φ x y) (h : ∀ (x : M) (y : N) (x' : M) (y' : N), cM x x'cN y y'HEq (f x y) (f x' y')) :
φ a b

A version of Quotient.hrecOn₂' for quotients by Con.

Equations
@[simp]
theorem AddCon.hrec_on₂_coe {M : Type u_1} {N : Type u_2} [Add M] [Add N] {cM : AddCon M} {cN : AddCon N} {φ : cM.QuotientcN.QuotientSort u_4} (a : M) (b : N) (f : (x : M) → (y : N) → φ x y) (h : ∀ (x : M) (y : N) (x' : M) (y' : N), cM x x'cN y y'HEq (f x y) (f x' y')) :
AddCon.hrecOn₂ (↑a) (↑b) f h = f a b
@[simp]
theorem Con.hrec_on₂_coe {M : Type u_1} {N : Type u_2} [Mul M] [Mul N] {cM : Con M} {cN : Con N} {φ : cM.QuotientcN.QuotientSort u_4} (a : M) (b : N) (f : (x : M) → (y : N) → φ x y) (h : ∀ (x : M) (y : N) (x' : M) (y' : N), cM x x'cN y y'HEq (f x y) (f x' y')) :
Con.hrecOn₂ (↑a) (↑b) f h = f a b
theorem AddCon.induction_on {M : Type u_1} [Add M] {c : AddCon M} {C : c.QuotientProp} (q : c.Quotient) (H : ∀ (x : M), C x) :
C q

The inductive principle used to prove propositions about the elements of a quotient by an additive congruence relation.

theorem Con.induction_on {M : Type u_1} [Mul M] {c : Con M} {C : c.QuotientProp} (q : c.Quotient) (H : ∀ (x : M), C x) :
C q

The inductive principle used to prove propositions about the elements of a quotient by a congruence relation.

theorem AddCon.induction_on₂ {M : Type u_1} {N : Type u_2} [Add M] [Add N] {c : AddCon M} {d : AddCon N} {C : c.Quotientd.QuotientProp} (p : c.Quotient) (q : d.Quotient) (H : ∀ (x : M) (y : N), C x y) :
C p q

A version of AddCon.induction_on for predicates which take two arguments.

theorem Con.induction_on₂ {M : Type u_1} {N : Type u_2} [Mul M] [Mul N] {c : Con M} {d : Con N} {C : c.Quotientd.QuotientProp} (p : c.Quotient) (q : d.Quotient) (H : ∀ (x : M) (y : N), C x y) :
C p q

A version of Con.induction_on for predicates which take two arguments.

@[simp]
theorem AddCon.eq {M : Type u_1} [Add M] (c : AddCon M) {a : M} {b : M} :
a = b c a b

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

@[simp]
theorem Con.eq {M : Type u_1} [Mul M] (c : Con M) {a : M} {b : M} :
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.

instance AddCon.hasAdd {M : Type u_1} [Add M] (c : AddCon M) :
Add c.Quotient

The addition induced on the quotient by an additive congruence relation on a type with an addition.

Equations
instance Con.hasMul {M : Type u_1} [Mul M] (c : Con M) :
Mul c.Quotient

The multiplication induced on the quotient by a congruence relation on a type with a multiplication.

Equations
@[simp]
theorem AddCon.add_ker_mk_eq {M : Type u_1} [Add M] (c : AddCon M) :
AddCon.addKer AddCon.toQuotient = c

The kernel of the quotient map induced by an additive congruence relation c equals c.

@[simp]
theorem Con.mul_ker_mk_eq {M : Type u_1} [Mul M] (c : Con M) :
Con.mulKer Con.toQuotient = c

The kernel of the quotient map induced by a congruence relation c equals c.

@[simp]
theorem AddCon.coe_add {M : Type u_1} [Add M] {c : AddCon M} (x : M) (y : M) :
(x + y) = x + y

The coercion to the quotient of an additive congruence relation commutes with addition (by definition).

@[simp]
theorem Con.coe_mul {M : Type u_1} [Mul M] {c : Con M} (x : M) (y : M) :
(x * y) = x * y

The coercion to the quotient of a congruence relation commutes with multiplication (by definition).

@[simp]
theorem AddCon.liftOn_coe {M : Type u_1} [Add M] {β : Sort u_4} (c : AddCon M) (f : Mβ) (h : ∀ (a b : M), c a bf a = f b) (x : M) :
AddCon.liftOn (↑x) f h = f x

Definition of the function on the quotient by an additive congruence relation c induced by a function that is constant on c's equivalence classes.

@[simp]
theorem Con.liftOn_coe {M : Type u_1} [Mul M] {β : Sort u_4} (c : Con M) (f : Mβ) (h : ∀ (a b : M), c a bf a = f b) (x : M) :
Con.liftOn (↑x) f h = f x

Definition of the function on the quotient by a congruence relation c induced by a function that is constant on c's equivalence classes.

def AddCon.congr {M : Type u_1} [Add M] {c : AddCon M} {d : AddCon M} (h : c = d) :
c.Quotient ≃+ d.Quotient

Makes an additive isomorphism of quotients by two additive congruence relations, given that the relations are equal.

Equations
theorem AddCon.congr.proof_1 {M : Type u_1} [Add M] {c : AddCon M} {d : AddCon M} (h : c = d) (x : M) (y : M) :
c x y d x y
theorem AddCon.congr.proof_2 {M : Type u_1} [Add M] {c : AddCon M} {d : AddCon M} (h : c = d) (x : c.Quotient) (y : c.Quotient) :
(Quotient.congr (Equiv.refl M) ).toFun (x + y) = (Quotient.congr (Equiv.refl M) ).toFun x + (Quotient.congr (Equiv.refl M) ).toFun y
def Con.congr {M : Type u_1} [Mul M] {c : Con M} {d : Con M} (h : c = d) :
c.Quotient ≃* d.Quotient

Makes an isomorphism of quotients by two congruence relations, given that the relations are equal.

Equations
instance AddCon.instLE {M : Type u_1} [Add M] :

For additive congruence relations c, d on a type M with an addition, c ≤ d iff ∀ x y ∈ M, x is related to y by d if x is related to y by c.

Equations
  • AddCon.instLE = { le := fun (c d : AddCon M) => ∀ ⦃x y : M⦄, c x yd x y }
instance Con.instLE {M : Type u_1} [Mul M] :
LE (Con M)

For congruence relations c, d on a type M with a multiplication, c ≤ d iff ∀ x y ∈ M, x is related to y by d if x is related to y by c.

Equations
  • Con.instLE = { le := fun (c d : Con M) => ∀ ⦃x y : M⦄, c x yd x y }
theorem AddCon.le_def {M : Type u_1} [Add M] {c : AddCon M} {d : AddCon M} :
c d ∀ {x y : M}, c x yd x y

Definition of for additive congruence relations.

theorem Con.le_def {M : Type u_1} [Mul M] {c : Con M} {d : Con M} :
c d ∀ {x y : M}, c x yd x y

Definition of for congruence relations.

theorem AddCon.instInfSet.proof_2 {M : Type u_1} [Add M] (S : Set (AddCon M)) :
∀ {w x y z : M}, Setoid.r w xSetoid.r y zcS, c (w + y) (x + z)
instance AddCon.instInfSet {M : Type u_1} [Add M] :

The infimum of a set of additive congruence relations on a given type with an addition.

Equations
  • AddCon.instInfSet = { sInf := fun (S : Set (AddCon M)) => { r := fun (x y : M) => cS, c x y, iseqv := , add' := } }
theorem AddCon.instInfSet.proof_1 {M : Type u_1} [Add M] (S : Set (AddCon M)) :
Equivalence fun (x y : M) => cS, c x y
instance Con.instInfSet {M : Type u_1} [Mul M] :

The infimum of a set of congruence relations on a given type with a multiplication.

Equations
  • Con.instInfSet = { sInf := fun (S : Set (Con M)) => { r := fun (x y : M) => cS, c x y, iseqv := , mul' := } }
theorem AddCon.sInf_toSetoid {M : Type u_1} [Add M] (S : Set (AddCon M)) :
(sInf S).toSetoid = sInf (AddCon.toSetoid '' S)

The infimum of a set of additive congruence relations is the same as the infimum of the set's image under the map to the underlying equivalence relation.

theorem Con.sInf_toSetoid {M : Type u_1} [Mul M] (S : Set (Con M)) :
(sInf S).toSetoid = sInf (Con.toSetoid '' S)

The infimum of a set of congruence relations is the same as the infimum of the set's image under the map to the underlying equivalence relation.

@[simp]
theorem AddCon.coe_sInf {M : Type u_1} [Add M] (S : Set (AddCon M)) :
(sInf S) = sInf (DFunLike.coe '' S)

The infimum of a set of additive congruence relations is the same as the infimum of the set's image under the map to the underlying binary relation.

@[simp]
theorem Con.coe_sInf {M : Type u_1} [Mul M] (S : Set (Con M)) :
(sInf S) = sInf (DFunLike.coe '' S)

The infimum of a set of congruence relations is the same as the infimum of the set's image under the map to the underlying binary relation.

@[simp]
theorem AddCon.coe_iInf {M : Type u_1} [Add M] {ι : Sort u_4} (f : ιAddCon M) :
(iInf f) = ⨅ (i : ι), (f i)
@[simp]
theorem Con.coe_iInf {M : Type u_1} [Mul M] {ι : Sort u_4} (f : ιCon M) :
(iInf f) = ⨅ (i : ι), (f i)
theorem AddCon.instPartialOrder.proof_3 {M : Type u_1} [Add M] :
∀ (a b : AddCon M), a < b a < b
instance AddCon.instPartialOrder {M : Type u_1} [Add M] :
Equations
theorem AddCon.instPartialOrder.proof_4 {M : Type u_1} [Add M] :
∀ (x x_1 : AddCon M), x x_1x_1 xx = x_1
theorem AddCon.instPartialOrder.proof_2 {M : Type u_1} [Add M] :
∀ (x x_1 x_2 : AddCon M), x x_1x_1 x_2∀ (x_3 x_4 : M), x x_3 x_4x_2 x_3 x_4
theorem AddCon.instPartialOrder.proof_1 {M : Type u_1} [Add M] :
∀ (x : AddCon M) (x_1 x_2 : M), x x_1 x_2x x_1 x_2
instance Con.instPartialOrder {M : Type u_1} [Mul M] :
Equations
theorem AddCon.instCompleteLattice.proof_10 {M : Type u_1} [Add M] :
∀ {w x y z : M}, Setoid.r w xSetoid.r y zSetoid.r (w + y) (x + z)
theorem AddCon.instCompleteLattice.proof_2 {M : Type u_1} [Add M] (c : AddCon M) (d : AddCon M) :
∀ {w x y z : M}, Setoid.r w xSetoid.r y zc.Rel (w + y) (x + z) d.Rel (w + y) (x + z)
theorem AddCon.instCompleteLattice.proof_3 {M : Type u_1} [Add M] :
∀ (x x_1 : AddCon M) (x_2 x_3 : M), (x x_1) x_2 x_3x.Rel x_2 x_3
theorem AddCon.instCompleteLattice.proof_4 {M : Type u_1} [Add M] :
∀ (x x_1 : AddCon M) (x_2 x_3 : M), (x x_1) x_2 x_3x_1.Rel x_2 x_3
theorem AddCon.instCompleteLattice.proof_7 {M : Type u_1} [Add M] (s : Set (AddCon M)) (a : AddCon M) :
(∀ bs, b a)sSup s a
theorem AddCon.instCompleteLattice.proof_8 {M : Type u_1} [Add M] (s : Set (AddCon M)) (a : AddCon M) :
a ssInf s a
theorem AddCon.instCompleteLattice.proof_6 {M : Type u_1} [Add M] (s : Set (AddCon M)) (a : AddCon M) :
a sa sSup s
theorem AddCon.instCompleteLattice.proof_5 {M : Type u_1} [Add M] :
∀ (x x_1 x_2 : AddCon M), x x_1x x_2∀ (x_3 x_4 : M), x x_3 x_4x_1.Rel x_3 x_4 x_2.Rel x_3 x_4
theorem AddCon.instCompleteLattice.proof_9 {M : Type u_1} [Add M] (s : Set (AddCon M)) (a : AddCon M) :
(∀ bs, a b)a sInf s
theorem AddCon.instCompleteLattice.proof_11 {M : Type u_1} [Add M] (c : AddCon M) (x : M) (y : M) (h : x y) :
c x y

The complete lattice of additive congruence relations on a given type with an addition.

Equations
instance Con.instCompleteLattice {M : Type u_1} [Mul M] :

The complete lattice of congruence relations on a given type with a multiplication.

Equations
@[simp]
theorem AddCon.coe_inf {M : Type u_1} [Add M] {c : AddCon M} {d : AddCon M} :
(c d) = c d

The infimum of two additive congruence relations equals the infimum of the underlying binary operations.

@[simp]
theorem Con.coe_inf {M : Type u_1} [Mul M] {c : Con M} {d : Con M} :
(c d) = c d

The infimum of two congruence relations equals the infimum of the underlying binary operations.

theorem AddCon.inf_iff_and {M : Type u_1} [Add M] {c : AddCon M} {d : AddCon M} {x : M} {y : M} :
(c d) x y c x y d x y

Definition of the infimum of two additive congruence relations.

theorem Con.inf_iff_and {M : Type u_1} [Mul M] {c : Con M} {d : Con M} {x : M} {y : M} :
(c d) x y c x y d x y

Definition of the infimum of two congruence relations.

theorem AddCon.addConGen_eq {M : Type u_1} [Add M] (r : MMProp) :
addConGen r = sInf {s : AddCon M | ∀ (x y : M), r x ys x y}

The inductively defined smallest additive congruence relation containing a binary relation r equals the infimum of the set of additive congruence relations containing r.

theorem Con.conGen_eq {M : Type u_1} [Mul M] (r : MMProp) :
conGen r = sInf {s : Con M | ∀ (x y : M), r x ys x y}

The inductively defined smallest congruence relation containing a binary relation r equals the infimum of the set of congruence relations containing r.

theorem AddCon.addConGen_le {M : Type u_1} [Add M] {r : MMProp} {c : AddCon M} (h : ∀ (x y : M), r x yc x y) :

The smallest additive congruence relation containing a binary relation r is contained in any additive congruence relation containing r.

theorem Con.conGen_le {M : Type u_1} [Mul M] {r : MMProp} {c : Con M} (h : ∀ (x y : M), r x yc x y) :

The smallest congruence relation containing a binary relation r is contained in any congruence relation containing r.

theorem AddCon.addConGen_mono {M : Type u_1} [Add M] {r : MMProp} {s : MMProp} (h : ∀ (x y : M), r x ys x y) :

Given binary relations r, s with r contained in s, the smallest additive congruence relation containing s contains the smallest additive congruence relation containing r.

theorem Con.conGen_mono {M : Type u_1} [Mul M] {r : MMProp} {s : MMProp} (h : ∀ (x y : M), r x ys x y) :

Given binary relations r, s with r contained in s, the smallest congruence relation containing s contains the smallest congruence relation containing r.

@[simp]
theorem AddCon.addConGen_of_addCon {M : Type u_1} [Add M] (c : AddCon M) :
addConGen c = c

Additive congruence relations equal the smallest additive congruence relation in which they are contained.

@[simp]
theorem Con.conGen_of_con {M : Type u_1} [Mul M] (c : Con M) :
conGen c = c

Congruence relations equal the smallest congruence relation in which they are contained.

theorem AddCon.addConGen_idem {M : Type u_1} [Add M] (r : MMProp) :

The map sending a binary relation to the smallest additive congruence relation in which it is contained is idempotent.

theorem Con.conGen_idem {M : Type u_1} [Mul M] (r : MMProp) :

The map sending a binary relation to the smallest congruence relation in which it is contained is idempotent.

theorem AddCon.sup_eq_addConGen {M : Type u_1} [Add M] (c : AddCon M) (d : AddCon M) :
c d = addConGen fun (x y : M) => c x y d x y

The supremum of additive congruence relations c, d equals the smallest additive congruence relation containing the binary relation 'x is related to y by c or d'.

theorem Con.sup_eq_conGen {M : Type u_1} [Mul M] (c : Con M) (d : Con M) :
c d = conGen fun (x y : M) => c x y d x y

The supremum of congruence relations c, d equals the smallest congruence relation containing the binary relation 'x is related to y by c or d'.

theorem AddCon.sup_def {M : Type u_1} [Add M] {c : AddCon M} {d : AddCon M} :
c d = addConGen (c d)

The supremum of two additive congruence relations equals the smallest additive congruence relation containing the supremum of the underlying binary operations.

theorem Con.sup_def {M : Type u_1} [Mul M] {c : Con M} {d : Con M} :
c d = conGen (c d)

The supremum of two congruence relations equals the smallest congruence relation containing the supremum of the underlying binary operations.

theorem AddCon.sSup_eq_addConGen {M : Type u_1} [Add M] (S : Set (AddCon M)) :
sSup S = addConGen fun (x y : M) => cS, c x y

The supremum of a set of additive congruence relations S equals the smallest additive congruence relation containing the binary relation 'there exists c ∈ S such that x is related to y by c'.

theorem Con.sSup_eq_conGen {M : Type u_1} [Mul M] (S : Set (Con M)) :
sSup S = conGen fun (x y : M) => cS, c x y

The supremum of a set of congruence relations S equals the smallest congruence relation containing the binary relation 'there exists c ∈ S such that x is related to y by c'.

theorem AddCon.sSup_def {M : Type u_1} [Add M] {S : Set (AddCon M)} :
sSup S = addConGen (sSup (DFunLike.coe '' S))

The supremum of a set of additive congruence relations is the same as the smallest additive congruence relation containing the supremum of the set's image under the map to the underlying binary relation.

theorem Con.sSup_def {M : Type u_1} [Mul M] {S : Set (Con M)} :
sSup S = conGen (sSup (DFunLike.coe '' S))

The supremum of a set of congruence relations is the same as the smallest congruence relation containing the supremum of the set's image under the map to the underlying binary relation.

theorem AddCon.gi.proof_1 (M : Type u_1) [Add M] :
∀ (x : MMProp) (c : AddCon M), addConGen x c x c
theorem AddCon.gi.proof_3 (M : Type u_1) [Add M] :
∀ (x : MMProp) (x_1 : (addConGen x) x), (fun (r : MMProp) (x : (addConGen r) r) => addConGen r) x x_1 = (fun (r : MMProp) (x : (addConGen r) r) => addConGen r) x x_1
theorem AddCon.gi.proof_2 (M : Type u_1) [Add M] (x : AddCon M) :
x addConGen x
def AddCon.gi (M : Type u_1) [Add M] :
GaloisInsertion addConGen DFunLike.coe

There is a Galois insertion of additive congruence relations on a type with an addition M into binary relations on M.

Equations
def Con.gi (M : Type u_1) [Mul M] :
GaloisInsertion conGen DFunLike.coe

There is a Galois insertion of congruence relations on a type with a multiplication M into binary relations on M.

Equations
  • Con.gi M = { choice := fun (r : MMProp) (x : (conGen r) r) => conGen r, gc := , le_l_u := , choice_eq := }
def AddCon.mapGen {M : Type u_1} {N : Type u_2} [Add M] [Add N] (c : AddCon M) (f : MN) :

Given a function f, the smallest additive congruence relation containing the binary relation on f's image defined by 'x ≈ y iff the elements of f⁻¹(x) are related to the elements of f⁻¹(y) by an additive congruence relation c.'

Equations
  • c.mapGen f = addConGen fun (x y : N) => ∃ (a : M) (b : M), f a = x f b = y c a b
def Con.mapGen {M : Type u_1} {N : Type u_2} [Mul M] [Mul N] (c : Con M) (f : MN) :
Con N

Given a function f, the smallest congruence relation containing the binary relation on f's image defined by 'x ≈ y iff the elements of f⁻¹(x) are related to the elements of f⁻¹(y) by a congruence relation c.'

Equations
  • c.mapGen f = conGen fun (x y : N) => ∃ (a : M) (b : M), f a = x f b = y c a b
def AddCon.mapOfSurjective {M : Type u_1} {N : Type u_2} [Add M] [Add N] (c : AddCon M) (f : MN) (H : ∀ (x y : M), f (x + y) = f x + f y) (h : AddCon.addKer f H c) (hf : Function.Surjective f) :

Given a surjective addition-preserving function f whose kernel is contained in an additive congruence relation c, the additive congruence relation on f's codomain defined by 'x ≈ y iff the elements of f⁻¹(x) are related to the elements of f⁻¹(y) by c.'

Equations
  • c.mapOfSurjective f H h hf = { toSetoid := c.mapOfSurjective f h hf, add' := }
theorem AddCon.mapOfSurjective.proof_1 {M : Type u_2} {N : Type u_1} [Add M] [Add N] (c : AddCon M) (f : MN) (H : ∀ (x y : M), f (x + y) = f x + f y) (h : AddCon.addKer f H c) (hf : Function.Surjective f) :
∀ {w x y z : N}, Setoid.r w xSetoid.r y zSetoid.r (w + y) (x + z)
def Con.mapOfSurjective {M : Type u_1} {N : Type u_2} [Mul M] [Mul N] (c : Con M) (f : MN) (H : ∀ (x y : M), f (x * y) = f x * f y) (h : Con.mulKer f H c) (hf : Function.Surjective f) :
Con N

Given a surjective multiplicative-preserving function f whose kernel is contained in a congruence relation c, the congruence relation on f's codomain defined by 'x ≈ y iff the elements of f⁻¹(x) are related to the elements of f⁻¹(y) by c.'

Equations
  • c.mapOfSurjective f H h hf = { toSetoid := c.mapOfSurjective f h hf, mul' := }
theorem AddCon.mapOfSurjective_eq_mapGen {M : Type u_1} {N : Type u_2} [Add M] [Add N] {c : AddCon M} {f : MN} (H : ∀ (x y : M), f (x + y) = f x + f y) (h : AddCon.addKer f H c) (hf : Function.Surjective f) :
c.mapGen f = c.mapOfSurjective f H h hf

A specialization of 'the smallest additive congruence relation containing an additive congruence relation c equals c'.

theorem Con.mapOfSurjective_eq_mapGen {M : Type u_1} {N : Type u_2} [Mul M] [Mul N] {c : Con M} {f : MN} (H : ∀ (x y : M), f (x * y) = f x * f y) (h : Con.mulKer f H c) (hf : Function.Surjective f) :
c.mapGen f = c.mapOfSurjective f H h hf

A specialization of 'the smallest congruence relation containing a congruence relation c equals c'.

theorem AddCon.comap.proof_1 {M : Type u_1} {N : Type u_2} [Add M] [Add N] (f : MN) (H : ∀ (x y : M), f (x + y) = f x + f y) (c : AddCon N) (w : M) (x : M) (y : M) (z : M) (h1 : Setoid.r w x) (h2 : Setoid.r y z) :
c (f (w + y)) (f (x + z))
def AddCon.comap {M : Type u_1} {N : Type u_2} [Add M] [Add N] (f : MN) (H : ∀ (x y : M), f (x + y) = f x + f y) (c : AddCon N) :

Given types with additions M, N and an additive congruence relation c on N, an addition-preserving map f : M → N induces an additive congruence relation on f's domain defined by 'x ≈ y iff f(x) is related to f(y) by c.'

Equations
def Con.comap {M : Type u_1} {N : Type u_2} [Mul M] [Mul N] (f : MN) (H : ∀ (x y : M), f (x * y) = f x * f y) (c : Con N) :
Con M

Given types with multiplications M, N and a congruence relation c on N, a multiplication-preserving map f : M → N induces a congruence relation on f's domain defined by 'x ≈ y iff f(x) is related to f(y) by c.'

Equations
@[simp]
theorem AddCon.comap_rel {M : Type u_1} {N : Type u_2} [Add M] [Add N] {f : MN} (H : ∀ (x y : M), f (x + y) = f x + f y) {c : AddCon N} {x : M} {y : M} :
(AddCon.comap f H c) x y c (f x) (f y)
@[simp]
theorem Con.comap_rel {M : Type u_1} {N : Type u_2} [Mul M] [Mul N] {f : MN} (H : ∀ (x y : M), f (x * y) = f x * f y) {c : Con N} {x : M} {y : M} :
(Con.comap f H c) x y c (f x) (f y)
theorem AddCon.correspondence.proof_7 {M : Type u_1} [Add M] (c : AddCon M) (s : { d : AddCon M // c d }) (t : { d : AddCon M // c d }) :
{ toFun := fun (d : { d : AddCon M // c d }) => (↑d).mapOfSurjective AddCon.toQuotient , invFun := fun (d : AddCon c.Quotient) => AddCon.comap AddCon.toQuotient d, , left_inv := , right_inv := } s { toFun := fun (d : { d : AddCon M // c d }) => (↑d).mapOfSurjective AddCon.toQuotient , invFun := fun (d : AddCon c.Quotient) => AddCon.comap AddCon.toQuotient d, , left_inv := , right_inv := } t s t
theorem AddCon.correspondence.proof_5 {M : Type u_1} [Add M] (c : AddCon M) (d : { d : AddCon M // c d }) :
(fun (d : AddCon c.Quotient) => AddCon.comap AddCon.toQuotient d, ) ((fun (d : { d : AddCon M // c d }) => (↑d).mapOfSurjective AddCon.toQuotient ) d) = d
theorem AddCon.correspondence.proof_4 {M : Type u_1} [Add M] (c : AddCon M) (d : AddCon c.Quotient) (x : M) (y : M) (h : c x y) :
d x y
theorem AddCon.correspondence.proof_2 {M : Type u_1} [Add M] (c : AddCon M) (d : { d : AddCon M // c d }) :
AddCon.addKer AddCon.toQuotient d
def AddCon.correspondence {M : Type u_1} [Add M] (c : AddCon M) :
{ d : AddCon M // c d } ≃o AddCon c.Quotient

Given an additive congruence relation c on a type M with an addition, the order-preserving bijection between the set of additive congruence relations containing c and the additive congruence relations on the quotient of M by c.

Equations
  • One or more equations did not get rendered due to their size.
theorem AddCon.correspondence.proof_3 {M : Type u_1} [Add M] (c : AddCon M) (q : Quotient c.toSetoid) :
∃ (a : M), a = q
theorem AddCon.correspondence.proof_1 {M : Type u_1} [Add M] (c : AddCon M) (x : M) (y : M) :
(x + y) = (x + y)
theorem AddCon.correspondence.proof_6 {M : Type u_1} [Add M] (c : AddCon M) (d : AddCon c.Quotient) :
(fun (d : { d : AddCon M // c d }) => (↑d).mapOfSurjective AddCon.toQuotient ) ((fun (d : AddCon c.Quotient) => AddCon.comap AddCon.toQuotient d, ) d) = d
def Con.correspondence {M : Type u_1} [Mul M] (c : Con M) :
{ d : Con M // c d } ≃o Con c.Quotient

Given a congruence relation c on a type M with a multiplication, the order-preserving bijection between the set of congruence relations containing c and the congruence relations on the quotient of M by c.

Equations
  • One or more equations did not get rendered due to their size.
theorem AddCon.addZeroClass.proof_2 {M : Type u_1} [AddZeroClass M] (c : AddCon M) (x : c.Quotient) :
x + 0 = x
instance AddCon.addZeroClass {M : Type u_1} [AddZeroClass M] (c : AddCon M) :
AddZeroClass c.Quotient

The quotient of an AddMonoid by an additive congruence relation is an AddMonoid.

Equations
theorem AddCon.addZeroClass.proof_1 {M : Type u_1} [AddZeroClass M] (c : AddCon M) (x : c.Quotient) :
0 + x = x
instance Con.mulOneClass {M : Type u_1} [MulOneClass M] (c : Con M) :
MulOneClass c.Quotient

The quotient of a monoid by a congruence relation is a monoid.

Equations
@[simp]
theorem AddCon.coe_zero {M : Type u_1} [AddZeroClass M] {c : AddCon M} :
0 = 0

The 0 of the quotient of an AddMonoid by an additive congruence relation is the equivalence class of the AddMonoid's 0.

@[simp]
theorem Con.coe_one {M : Type u_1} [MulOneClass M] {c : Con M} :
1 = 1

The 1 of the quotient of a monoid by a congruence relation is the equivalence class of the monoid's 1.

def AddCon.addSubmonoid {M : Type u_1} [AddZeroClass M] (c : AddCon M) :

The AddSubmonoid of M × M defined by an additive congruence relation on an AddMonoid M.

Equations
  • c = { carrier := {x : M × M | c x.1 x.2}, add_mem' := , zero_mem' := }
theorem AddCon.addSubmonoid.proof_1 {M : Type u_1} [AddZeroClass M] (c : AddCon M) :
∀ {a b : M × M}, c a.1 a.2c b.1 b.2c (a.1 + b.1) (a.2 + b.2)
def Con.submonoid {M : Type u_1} [MulOneClass M] (c : Con M) :

The submonoid of M × M defined by a congruence relation on a monoid M.

Equations
  • c = { carrier := {x : M × M | c x.1 x.2}, mul_mem' := , one_mem' := }
def AddCon.ofAddSubmonoid {M : Type u_1} [AddZeroClass M] (N : AddSubmonoid (M × M)) (H : Equivalence fun (x y : M) => (x, y) N) :

The additive congruence relation on an AddMonoid M from an AddSubmonoid of M × M for which membership is an equivalence relation.

Equations
theorem AddCon.ofAddSubmonoid.proof_1 {M : Type u_1} [AddZeroClass M] (N : AddSubmonoid (M × M)) :
∀ {w x y z : M}, (w, x) N(y, z) N(w, x) + (y, z) N
def Con.ofSubmonoid {M : Type u_1} [MulOneClass M] (N : Submonoid (M × M)) (H : Equivalence fun (x y : M) => (x, y) N) :
Con M

The congruence relation on a monoid M from a submonoid of M × M for which membership is an equivalence relation.

Equations
instance AddCon.toAddSubmonoid {M : Type u_1} [AddZeroClass M] :

Coercion from a congruence relation c on an AddMonoid M to the AddSubmonoid of M × M whose elements are (x, y) such that x is related to y by c.

Equations
  • AddCon.toAddSubmonoid = { coe := fun (c : AddCon M) => c }
instance Con.toSubmonoid {M : Type u_1} [MulOneClass M] :
Coe (Con M) (Submonoid (M × M))

Coercion from a congruence relation c on a monoid M to the submonoid of M × M whose elements are (x, y) such that x is related to y by c.

Equations
  • Con.toSubmonoid = { coe := fun (c : Con M) => c }
theorem AddCon.mem_coe {M : Type u_1} [AddZeroClass M] {c : AddCon M} {x : M} {y : M} :
(x, y) c (x, y) c
theorem Con.mem_coe {M : Type u_1} [MulOneClass M] {c : Con M} {x : M} {y : M} :
(x, y) c (x, y) c
theorem AddCon.to_addSubmonoid_inj {M : Type u_1} [AddZeroClass M] (c : AddCon M) (d : AddCon M) (H : c = d) :
c = d
theorem Con.to_submonoid_inj {M : Type u_1} [MulOneClass M] (c : Con M) (d : Con M) (H : c = d) :
c = d
theorem AddCon.le_iff {M : Type u_1} [AddZeroClass M] {c : AddCon M} {d : AddCon M} :
c d c d
theorem Con.le_iff {M : Type u_1} [MulOneClass M] {c : Con M} {d : Con M} :
c d c d
def AddCon.ker {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) :

The kernel of an AddMonoid homomorphism as an additive congruence relation.

Equations
theorem AddCon.ker.proof_1 {M : Type u_2} {P : Type u_1} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) (x : M) (y : M) :
f (x + y) = f x + f y
def Con.ker {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] (f : M →* P) :
Con M

The kernel of a monoid homomorphism as a congruence relation.

Equations
@[simp]
theorem AddCon.ker_rel {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) {x : M} {y : M} :
(AddCon.ker f) x y f x = f y

The definition of the additive congruence relation defined by an AddMonoid homomorphism's kernel.

@[simp]
theorem Con.ker_rel {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] (f : M →* P) {x : M} {y : M} :
(Con.ker f) x y f x = f y

The definition of the congruence relation defined by a monoid homomorphism's kernel.

instance AddCon.Quotient.inhabited {M : Type u_1} [AddZeroClass M] {c : AddCon M} :
Inhabited c.Quotient

There exists an element of the quotient of an AddMonoid by a congruence relation (namely 0).

Equations
  • AddCon.Quotient.inhabited = { default := 0 }
instance Con.Quotient.inhabited {M : Type u_1} [MulOneClass M] {c : Con M} :
Inhabited c.Quotient

There exists an element of the quotient of a monoid by a congruence relation (namely 1).

Equations
  • Con.Quotient.inhabited = { default := 1 }
theorem AddCon.mk'.proof_2 {M : Type u_1} [AddZeroClass M] (c : AddCon M) :
∀ (x x_1 : M), { toFun := AddCon.toQuotient, map_zero' := }.toFun (x + x_1) = { toFun := AddCon.toQuotient, map_zero' := }.toFun (x + x_1)
theorem AddCon.mk'.proof_1 {M : Type u_1} [AddZeroClass M] (c : AddCon M) :
0 = 0
def AddCon.mk' {M : Type u_1} [AddZeroClass M] (c : AddCon M) :
M →+ c.Quotient

The natural homomorphism from an AddMonoid to its quotient by an additive congruence relation.

Equations
  • c.mk' = { toFun := AddCon.toQuotient, map_zero' := , map_add' := }
def Con.mk' {M : Type u_1} [MulOneClass M] (c : Con M) :
M →* c.Quotient

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

Equations
  • c.mk' = { toFun := Con.toQuotient, map_one' := , map_mul' := }
@[simp]
theorem AddCon.mk'_ker {M : Type u_1} [AddZeroClass M] (c : AddCon M) :
AddCon.ker c.mk' = c

The kernel of the natural homomorphism from an AddMonoid to its quotient by an additive congruence relation c equals c.

@[simp]
theorem Con.mk'_ker {M : Type u_1} [MulOneClass M] (c : Con M) :
Con.ker c.mk' = c

The kernel of the natural homomorphism from a monoid to its quotient by a congruence relation c equals c.

theorem AddCon.mk'_surjective {M : Type u_1} [AddZeroClass M] {c : AddCon M} :

The natural homomorphism from an AddMonoid to its quotient by a congruence relation is surjective.

theorem Con.mk'_surjective {M : Type u_1} [MulOneClass M] {c : Con M} :

The natural homomorphism from a monoid to its quotient by a congruence relation is surjective.

@[simp]
theorem AddCon.coe_mk' {M : Type u_1} [AddZeroClass M] {c : AddCon M} :
c.mk' = AddCon.toQuotient
@[simp]
theorem Con.coe_mk' {M : Type u_1} [MulOneClass M] {c : Con M} :
c.mk' = Con.toQuotient
@[simp]
theorem AddCon.mrange_mk' {M : Type u_1} [AddZeroClass M] {c : AddCon M} :
@[simp]
theorem Con.mrange_mk' {M : Type u_1} [MulOneClass M] {c : Con M} :
theorem AddCon.ker_apply {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] {f : M →+ P} {x : M} {y : M} :
(AddCon.ker f) x y f x = f y
theorem Con.ker_apply {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] {f : M →* P} {x : M} {y : M} :
(Con.ker f) x y f x = f y
theorem AddCon.comap_eq {M : Type u_1} {N : Type u_2} [AddZeroClass M] [AddZeroClass N] {c : AddCon M} {f : N →+ M} :
AddCon.comap f c = AddCon.ker (c.mk'.comp f)

Given an AddMonoid homomorphism f : N → M and an additive congruence relation c on M, the additive congruence relation induced on N by f equals the kernel of c's quotient homomorphism composed with f.

theorem Con.comap_eq {M : Type u_1} {N : Type u_2} [MulOneClass M] [MulOneClass N] {c : Con M} {f : N →* M} :
Con.comap f c = Con.ker (c.mk'.comp f)

Given a monoid homomorphism f : N → M and a congruence relation c on M, the congruence relation induced on N by f equals the kernel of c's quotient homomorphism composed with f.

def AddCon.lift {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] (c : AddCon M) (f : M →+ P) (H : c AddCon.ker f) :
c.Quotient →+ P

The homomorphism on the quotient of an AddMonoid by an additive congruence relation c induced by a homomorphism constant on c's equivalence classes.

Equations
  • c.lift f H = { toFun := fun (x : c.Quotient) => AddCon.liftOn x f , map_zero' := , map_add' := }
theorem AddCon.lift.proof_1 {M : Type u_2} {P : Type u_1} [AddZeroClass M] [AddZeroClass P] (c : AddCon M) (f : M →+ P) (H : c AddCon.ker f) :
(fun (x : c.Quotient) => AddCon.liftOn x f ) 0 = 0
theorem AddCon.lift.proof_2 {M : Type u_1} {P : Type u_2} [AddZeroClass M] [AddZeroClass P] (c : AddCon M) (f : M →+ P) (H : c AddCon.ker f) (x : c.Quotient) (y : c.Quotient) :
{ toFun := fun (x : c.Quotient) => AddCon.liftOn x f , map_zero' := }.toFun (x + y) = { toFun := fun (x : c.Quotient) => AddCon.liftOn x f , map_zero' := }.toFun x + { toFun := fun (x : c.Quotient) => AddCon.liftOn x f , map_zero' := }.toFun y
def Con.lift {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] (c : Con M) (f : M →* P) (H : c Con.ker f) :
c.Quotient →* P

The homomorphism on the quotient of a monoid by a congruence relation c induced by a homomorphism constant on c's equivalence classes.

Equations
  • c.lift f H = { toFun := fun (x : c.Quotient) => Con.liftOn x f , map_one' := , map_mul' := }
theorem AddCon.lift_mk' {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] {c : AddCon M} {f : M →+ P} (H : c AddCon.ker f) (x : M) :
(c.lift f H) (c.mk' x) = f x

The diagram describing the universal property for quotients of AddMonoids commutes.

theorem Con.lift_mk' {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] {c : Con M} {f : M →* P} (H : c Con.ker f) (x : M) :
(c.lift f H) (c.mk' x) = f x

The diagram describing the universal property for quotients of monoids commutes.

@[simp]
theorem AddCon.lift_coe {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] {c : AddCon M} {f : M →+ P} (H : c AddCon.ker f) (x : M) :
(c.lift f H) x = f x

The diagram describing the universal property for quotients of AddMonoids commutes.

@[simp]
theorem Con.lift_coe {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] {c : Con M} {f : M →* P} (H : c Con.ker f) (x : M) :
(c.lift f H) x = f x

The diagram describing the universal property for quotients of monoids commutes.

@[simp]
theorem AddCon.lift_comp_mk' {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] {c : AddCon M} {f : M →+ P} (H : c AddCon.ker f) :
(c.lift f H).comp c.mk' = f

The diagram describing the universal property for quotients of AddMonoids commutes.

@[simp]
theorem Con.lift_comp_mk' {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] {c : Con M} {f : M →* P} (H : c Con.ker f) :
(c.lift f H).comp c.mk' = f

The diagram describing the universal property for quotients of monoids commutes.

@[simp]
theorem AddCon.lift_apply_mk' {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] {c : AddCon M} (f : c.Quotient →+ P) :
c.lift (f.comp c.mk') = f

Given a homomorphism f from the quotient of an AddMonoid by an additive congruence relation, f equals the homomorphism on the quotient induced by f composed with the natural map from the AddMonoid to the quotient.

@[simp]
theorem Con.lift_apply_mk' {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] {c : Con M} (f : c.Quotient →* P) :
c.lift (f.comp c.mk') = f

Given a homomorphism f from the quotient of a monoid by a congruence relation, f equals the homomorphism on the quotient induced by f composed with the natural map from the monoid to the quotient.

theorem AddCon.lift_funext {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] {c : AddCon M} (f : c.Quotient →+ P) (g : c.Quotient →+ P) (h : ∀ (a : M), f a = g a) :
f = g

Homomorphisms on the quotient of an AddMonoid by an additive congruence relation are equal if they are equal on elements that are coercions from the AddMonoid.

theorem Con.lift_funext {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] {c : Con M} (f : c.Quotient →* P) (g : c.Quotient →* P) (h : ∀ (a : M), f a = g a) :
f = g

Homomorphisms on the quotient of a monoid by a congruence relation are equal if they are equal on elements that are coercions from the monoid.

theorem AddCon.lift_unique {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] {c : AddCon M} {f : M →+ P} (H : c AddCon.ker f) (g : c.Quotient →+ P) (Hg : g.comp c.mk' = f) :
g = c.lift f H

The uniqueness part of the universal property for quotients of AddMonoids.

theorem Con.lift_unique {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] {c : Con M} {f : M →* P} (H : c Con.ker f) (g : c.Quotient →* P) (Hg : g.comp c.mk' = f) :
g = c.lift f H

The uniqueness part of the universal property for quotients of monoids.

theorem AddCon.lift_range {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] {c : AddCon M} {f : M →+ P} (H : c AddCon.ker f) :

Given an additive congruence relation c on an AddMonoid and a homomorphism f constant on c's equivalence classes, f has the same image as the homomorphism that f induces on the quotient.

theorem Con.lift_range {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] {c : Con M} {f : M →* P} (H : c Con.ker f) :

Given a congruence relation c on a monoid and a homomorphism f constant on c's equivalence classes, f has the same image as the homomorphism that f induces on the quotient.

theorem AddCon.lift_surjective_of_surjective {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] {c : AddCon M} {f : M →+ P} (h : c AddCon.ker f) (hf : Function.Surjective f) :
Function.Surjective (c.lift f h)

Surjective AddMonoid homomorphisms constant on an additive congruence relation c's equivalence classes induce a surjective homomorphism on c's quotient.

theorem Con.lift_surjective_of_surjective {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] {c : Con M} {f : M →* P} (h : c Con.ker f) (hf : Function.Surjective f) :
Function.Surjective (c.lift f h)

Surjective monoid homomorphisms constant on a congruence relation c's equivalence classes induce a surjective homomorphism on c's quotient.

theorem AddCon.ker_eq_lift_of_injective {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] (c : AddCon M) (f : M →+ P) (H : c AddCon.ker f) (h : Function.Injective (c.lift f H)) :

Given an AddMonoid homomorphism f from M to P, the kernel of f is the unique additive congruence relation on M whose induced map from the quotient of M to P is injective.

theorem Con.ker_eq_lift_of_injective {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] (c : Con M) (f : M →* P) (H : c Con.ker f) (h : Function.Injective (c.lift f H)) :

Given a monoid homomorphism f from M to P, the kernel of f is the unique congruence relation on M whose induced map from the quotient of M to P is injective.

theorem AddCon.kerLift.proof_1 {M : Type u_1} {P : Type u_2} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) :
∀ (x x_1 : M), (AddCon.ker f) x x_1(AddCon.ker f) x x_1
def AddCon.kerLift {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) :
(AddCon.ker f).Quotient →+ P

The homomorphism induced on the quotient of an AddMonoid by the kernel of an AddMonoid homomorphism.

Equations
def Con.kerLift {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] (f : M →* P) :
(Con.ker f).Quotient →* P

The homomorphism induced on the quotient of a monoid by the kernel of a monoid homomorphism.

Equations
@[simp]
theorem AddCon.kerLift_mk {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] {f : M →+ P} (x : M) :
(AddCon.kerLift f) x = f x

The diagram described by the universal property for quotients of AddMonoids, when the additive congruence relation is the kernel of the homomorphism, commutes.

@[simp]
theorem Con.kerLift_mk {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] {f : M →* P} (x : M) :
(Con.kerLift f) x = f x

The diagram described by the universal property for quotients of monoids, when the congruence relation is the kernel of the homomorphism, commutes.

@[simp]

Given an AddMonoid homomorphism f, the induced homomorphism on the quotient by f's kernel has the same image as f.

@[simp]

Given a monoid homomorphism f, the induced homomorphism on the quotient by f's kernel has the same image as f.

An AddMonoid homomorphism f induces an injective homomorphism on the quotient by f's kernel.

theorem Con.kerLift_injective {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] (f : M →* P) :

A monoid homomorphism f induces an injective homomorphism on the quotient by f's kernel.

theorem AddCon.map.proof_1 {M : Type u_1} [AddZeroClass M] (c : AddCon M) (d : AddCon M) (h : c d) (x : M) (y : M) (hc : c x y) :
(AddCon.ker d.mk') x y
def AddCon.map {M : Type u_1} [AddZeroClass M] (c : AddCon M) (d : AddCon M) (h : c d) :
c.Quotient →+ d.Quotient

Given additive congruence relations c, d on an AddMonoid such that d contains c, d's quotient map induces a homomorphism from the quotient by c to the quotient by d.

Equations
  • c.map d h = c.lift d.mk'
def Con.map {M : Type u_1} [MulOneClass M] (c : Con M) (d : Con M) (h : c d) :
c.Quotient →* d.Quotient

Given congruence relations c, d on a monoid such that d contains c, d's quotient map induces a homomorphism from the quotient by c to the quotient by d.

Equations
  • c.map d h = c.lift d.mk'
theorem AddCon.map_apply {M : Type u_1} [AddZeroClass M] {c : AddCon M} {d : AddCon M} (h : c d) (x : c.Quotient) :
(c.map d h) x = (c.lift d.mk' ) x

Given additive congruence relations c, d on an AddMonoid such that d contains c, the definition of the homomorphism from the quotient by c to the quotient by d induced by d's quotient map.

theorem Con.map_apply {M : Type u_1} [MulOneClass M] {c : Con M} {d : Con M} (h : c d) (x : c.Quotient) :
(c.map d h) x = (c.lift d.mk' ) x

Given congruence relations c, d on a monoid such that d contains c, the definition of the homomorphism from the quotient by c to the quotient by d induced by d's quotient map.

theorem AddCon.quotientKerEquivRange.proof_3 {M : Type u_1} {P : Type u_2} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) (a : (AddCon.ker f).Quotient) (b : (AddCon.ker f).Quotient) :
((AddEquiv.addSubmonoidCongr ).toAddMonoidHom.comp (AddCon.kerLift f).mrangeRestrict) (a + b) = ((AddEquiv.addSubmonoidCongr ).toAddMonoidHom.comp (AddCon.kerLift f).mrangeRestrict) a + ((AddEquiv.addSubmonoidCongr ).toAddMonoidHom.comp (AddCon.kerLift f).mrangeRestrict) b
theorem AddCon.quotientKerEquivRange.proof_1 {M : Type u_2} {P : Type u_1} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) :
AddMonoidHomClass ((AddCon.ker f).Quotient →+ P) (AddCon.ker f).Quotient P
noncomputable def AddCon.quotientKerEquivRange {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) :
(AddCon.ker f).Quotient ≃+ (AddMonoidHom.mrange f)

The first isomorphism theorem for AddMonoids.

Equations
theorem AddCon.quotientKerEquivRange.proof_2 {M : Type u_1} {P : Type u_2} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) :
Function.Bijective ((AddEquiv.addSubmonoidCongr ).toEquiv (AddCon.kerLift f).mrangeRestrict)
noncomputable def Con.quotientKerEquivRange {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] (f : M →* P) :
(Con.ker f).Quotient ≃* (MonoidHom.mrange f)

The first isomorphism theorem for monoids.

Equations
def AddCon.quotientKerEquivOfRightInverse {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) (g : PM) (hf : Function.RightInverse g f) :
(AddCon.ker f).Quotient ≃+ P

The first isomorphism theorem for AddMonoids in the case of a homomorphism with right inverse.

Equations
theorem AddCon.quotientKerEquivOfRightInverse.proof_1 {M : Type u_1} {P : Type u_2} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) (g : PM) (hf : Function.RightInverse g f) (x : (AddCon.ker f).Quotient) :
(AddCon.toQuotient g) ((AddCon.kerLift f) x) = x
theorem AddCon.quotientKerEquivOfRightInverse.proof_3 {M : Type u_1} {P : Type u_2} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) (x : (AddCon.ker f).Quotient) (y : (AddCon.ker f).Quotient) :
(↑(AddCon.kerLift f)).toFun (x + y) = (↑(AddCon.kerLift f)).toFun x + (↑(AddCon.kerLift f)).toFun y
theorem AddCon.quotientKerEquivOfRightInverse.proof_2 {M : Type u_2} {P : Type u_1} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) (g : PM) (hf : Function.RightInverse g f) (x : P) :
(AddCon.kerLift f) ((AddCon.toQuotient g) x) = x
@[simp]
theorem Con.quotientKerEquivOfRightInverse_symm_apply {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] (f : M →* P) (g : PM) (hf : Function.RightInverse g f) :
∀ (a : P), (Con.quotientKerEquivOfRightInverse f g hf).symm a = (Con.toQuotient g) a
@[simp]
theorem AddCon.quotientKerEquivOfRightInverse_symm_apply {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) (g : PM) (hf : Function.RightInverse g f) :
∀ (a : P), (AddCon.quotientKerEquivOfRightInverse f g hf).symm a = (AddCon.toQuotient g) a
@[simp]
theorem Con.quotientKerEquivOfRightInverse_apply {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] (f : M →* P) (g : PM) (hf : Function.RightInverse g f) (a : (Con.ker f).Quotient) :
@[simp]
theorem AddCon.quotientKerEquivOfRightInverse_apply {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) (g : PM) (hf : Function.RightInverse g f) (a : (AddCon.ker f).Quotient) :
def Con.quotientKerEquivOfRightInverse {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] (f : M →* P) (g : PM) (hf : Function.RightInverse g f) :
(Con.ker f).Quotient ≃* P

The first isomorphism theorem for monoids in the case of a homomorphism with right inverse.

Equations
noncomputable def AddCon.quotientKerEquivOfSurjective {M : Type u_1} {P : Type u_3} [AddZeroClass M] [AddZeroClass P] (f : M →+ P) (hf : Function.Surjective f) :
(AddCon.ker f).Quotient ≃+ P

The first isomorphism theorem for AddMonoids in the case of a surjective homomorphism.

For a computable version, see AddCon.quotientKerEquivOfRightInverse.

Equations
noncomputable def Con.quotientKerEquivOfSurjective {M : Type u_1} {P : Type u_3} [MulOneClass M] [MulOneClass P] (f : M →* P) (hf : Function.Surjective f) :
(Con.ker f).Quotient ≃* P

The first isomorphism theorem for Monoids in the case of a surjective homomorphism.

For a computable version, see Con.quotientKerEquivOfRightInverse.

Equations
noncomputable def AddCon.comapQuotientEquiv {M : Type u_1} {N : Type u_2} [AddZeroClass M] [AddZeroClass N] (c : AddCon M) (f : N →+ M) :
(AddCon.comap f c).Quotient ≃+ (AddMonoidHom.mrange (c.mk'.comp f))

The second isomorphism theorem for AddMonoids.

Equations
theorem AddCon.comapQuotientEquiv.proof_1 {M : Type u_1} {N : Type u_2} [AddZeroClass M] [AddZeroClass N] (c : AddCon M) :
AddMonoidHomClass (N →+ c.Quotient) N c.Quotient
noncomputable def Con.comapQuotientEquiv {M : Type u_1} {N : Type u_2} [MulOneClass M] [MulOneClass N] (c : Con M) (f : N →* M) :
(Con.comap f c).Quotient ≃* (MonoidHom.mrange (c.mk'.comp f))

The second isomorphism theorem for monoids.

Equations
theorem AddCon.quotientQuotientEquivQuotient.proof_1 {M : Type u_1} [AddZeroClass M] (c : AddCon M) (d : AddCon M) (h : c d) (x : (AddCon.ker (c.map d h)).Quotient) (y : (AddCon.ker (c.map d h)).Quotient) :
(c.quotientQuotientEquivQuotient d.toSetoid h).toFun (x + y) = (c.quotientQuotientEquivQuotient d.toSetoid h).toFun x + (c.quotientQuotientEquivQuotient d.toSetoid h).toFun y
def AddCon.quotientQuotientEquivQuotient {M : Type u_1} [AddZeroClass M] (c : AddCon M) (d : AddCon M) (h : c d) :
(AddCon.ker (c.map d h)).Quotient ≃+ d.Quotient

The third isomorphism theorem for AddMonoids.

Equations
  • c.quotientQuotientEquivQuotient d h = { toEquiv := c.quotientQuotientEquivQuotient d.toSetoid h, map_add' := }
def Con.quotientQuotientEquivQuotient {M : Type u_1} [MulOneClass M] (c : Con M) (d : Con M) (h : c d) :
(Con.ker (c.map d h)).Quotient ≃* d.Quotient

The third isomorphism theorem for monoids.

Equations
  • c.quotientQuotientEquivQuotient d h = { toEquiv := c.quotientQuotientEquivQuotient d.toSetoid h, map_mul' := }
theorem AddCon.nsmul {M : Type u_4} [AddMonoid M] (c : AddCon M) (n : ) {w : M} {x : M} :
c w xc (n w) (n x)

Additive congruence relations preserve natural scaling.

theorem Con.pow {M : Type u_4} [Monoid M] (c : Con M) (n : ) {w : M} {x : M} :
c w xc (w ^ n) (x ^ n)

Multiplicative congruence relations preserve natural powers.

instance AddCon.zero {M : Type u_1} [AddZeroClass M] (c : AddCon M) :
Zero c.Quotient
Equations
instance Con.one {M : Type u_1} [MulOneClass M] (c : Con M) :
One c.Quotient
Equations
theorem AddCon.vadd {α : Type u_4} {M : Type u_5} [AddZeroClass M] [VAdd α M] [VAddAssocClass α M M] (c : AddCon M) (a : α) {w : M} {x : M} (h : c w x) :
c (a +ᵥ w) (a +ᵥ x)
theorem Con.smul {α : Type u_4} {M : Type u_5} [MulOneClass M] [SMul α M] [IsScalarTower α M M] (c : Con M) (a : α) {w : M} {x : M} (h : c w x) :
c (a w) (a x)
instance AddCon.Quotient.nsmul {M : Type u_4} [AddMonoid M] (c : AddCon M) :
SMul c.Quotient
Equations
instance Con.instPowQuotientNat {M : Type u_4} [Monoid M] (c : Con M) :
Pow c.Quotient
Equations
  • c.instPowQuotientNat = { pow := fun (x : c.Quotient) (n : ) => Quotient.map' (fun (x : M) => x ^ n) x }
instance AddCon.addSemigroup {M : Type u_4} [AddSemigroup M] (c : AddCon M) :
AddSemigroup c.Quotient

The quotient of an AddSemigroup by an additive congruence relation is an AddSemigroup.

Equations
theorem AddCon.addSemigroup.proof_3 {M : Type u_1} [AddSemigroup M] (c : AddCon M) (a : c✝.Quotient) (b : c✝.Quotient) (c : c✝.Quotient) :
a + b + c = a + (b + c)
theorem AddCon.addSemigroup.proof_2 {M : Type u_1} [AddSemigroup M] (c : AddCon M) :
∀ (x x_1 : M), Quotient.mk'' (x + x_1) = Quotient.mk'' (x + x_1)
instance Con.semigroup {M : Type u_4} [Semigroup M] (c : Con M) :
Semigroup c.Quotient

The quotient of a semigroup by a congruence relation is a semigroup.

Equations
theorem AddCon.addCommSemigroup.proof_2 {M : Type u_1} [AddCommSemigroup M] (c : AddCon M) :
∀ (x x_1 : M), Quotient.mk'' (x + x_1) = Quotient.mk'' (x + x_1)
theorem AddCon.addCommSemigroup.proof_3 {M : Type u_1} [AddCommSemigroup M] (c : AddCon M) (a : c.Quotient) (b : c.Quotient) :
a + b = b + a
instance AddCon.addCommSemigroup {M : Type u_4} [AddCommSemigroup M] (c : AddCon M) :
AddCommSemigroup c.Quotient

The quotient of an AddCommSemigroup by an additive congruence relation is an AddCommSemigroup.

Equations
instance Con.commSemigroup {M : Type u_4} [CommSemigroup M] (c : Con M) :
CommSemigroup c.Quotient

The quotient of a commutative semigroup by a congruence relation is a semigroup.

Equations
instance AddCon.addMonoid {M : Type u_4} [AddMonoid M] (c : AddCon M) :
AddMonoid c.Quotient

The quotient of an AddMonoid by an additive congruence relation is an AddMonoid.

Equations
theorem AddCon.addMonoid.proof_5 {M : Type u_1} [AddMonoid M] (c : AddCon M) (a : c.Quotient) :
0 + a = a
theorem AddCon.addMonoid.proof_3 {M : Type u_1} [AddMonoid M] (c : AddCon M) :
∀ (x x_1 : M), Quotient.mk'' (x + x_1) = Quotient.mk'' (x + x_1)
theorem AddCon.addMonoid.proof_7 {M : Type u_1} [AddMonoid M] (c : AddCon M) (x : c.Quotient) :
theorem AddCon.addMonoid.proof_4 {M : Type u_1} [AddMonoid M] (c : AddCon M) :
∀ (x : M) (x_1 : ), Quotient.mk'' (x_1 x) = Quotient.mk'' (x_1 x)
theorem AddCon.addMonoid.proof_6 {M : Type u_1} [AddMonoid M] (c : AddCon M) (a : c.Quotient) :
a + 0 = a
theorem AddCon.addMonoid.proof_1 {M : Type u_1} [AddMonoid M] (c : AddCon M) :
Function.Surjective Quotient.mk''
theorem AddCon.addMonoid.proof_8 {M : Type u_1} [AddMonoid M] (c : AddCon M) (n : ) (x : c.Quotient) :
instance Con.monoid {M : Type u_4} [Monoid M] (c : Con M) :
Monoid c.Quotient

The quotient of a monoid by a congruence relation is a monoid.

Equations
theorem AddCon.addCommMonoid.proof_5 {M : Type u_1} [AddCommMonoid M] (c : AddCon M) (a : c.Quotient) (b : c.Quotient) :
a + b = b + a
theorem AddCon.addCommMonoid.proof_3 {M : Type u_1} [AddCommMonoid M] (c : AddCon M) :
∀ (x x_1 : M), Quotient.mk'' (x + x_1) = Quotient.mk'' (x + x_1)
theorem AddCon.addCommMonoid.proof_4 {M : Type u_1} [AddCommMonoid M] (c : AddCon M) :
∀ (x : M) (x_1 : ), Quotient.mk'' (x_1 x) = Quotient.mk'' (x_1 x)
instance AddCon.addCommMonoid {M : Type u_4} [AddCommMonoid M] (c : AddCon M) :
AddCommMonoid c.Quotient

The quotient of an AddCommMonoid by an additive congruence relation is an AddCommMonoid.

Equations
instance Con.commMonoid {M : Type u_4} [CommMonoid M] (c : Con M) :
CommMonoid c.Quotient

The quotient of a CommMonoid by a congruence relation is a CommMonoid.

Equations
theorem AddCon.map_of_add_left_rel_zero {M : Type u_1} [AddMonoid M] (c : AddCon M) (f : MM) (hf : ∀ (x : M), c (f x + x) 0) {x : M} {y : M} (h : c x y) :
c (f x) (f y)

Sometimes, an additive group is defined as a quotient of a monoid by an additive congruence relation. Usually, the inverse operation is defined as Setoid.map f _ for some f. This lemma allows to avoid code duplication in the definition of the inverse operation: instead of proving both ∀ x y, c x y → c (f x) (f y) (to define the operation) and ∀ x, c (f x + x) 0 (to prove the group laws), one can only prove the latter.

theorem Con.map_of_mul_left_rel_one {M : Type u_1} [Monoid M] (c : Con M) (f : MM) (hf : ∀ (x : M), c (f x * x) 1) {x : M} {y : M} (h : c x y) :
c (f x) (f y)

Sometimes, a group is defined as a quotient of a monoid by a congruence relation. Usually, the inverse operation is defined as Setoid.map f _ for some f. This lemma allows to avoid code duplication in the definition of the inverse operation: instead of proving both ∀ x y, c x y → c (f x) (f y) (to define the operation) and ∀ x, c (f x * x) 1 (to prove the group laws), one can only prove the latter.

theorem AddCon.neg {M : Type u_1} [AddGroup M] (c : AddCon M) {x : M} {y : M} (h : c x y) :
c (-x) (-y)

Additive congruence relations preserve negation.

theorem Con.inv {M : Type u_1} [Group M] (c : Con M) {x : M} {y : M} (h : c x y) :

Multiplicative congruence relations preserve inversion.

theorem AddCon.sub {M : Type u_1} [AddGroup M] (c : AddCon M) {w : M} {x : M} {y : M} {z : M} :
c w xc y zc (w - y) (x - z)

Additive congruence relations preserve subtraction.

theorem Con.div {M : Type u_1} [Group M] (c : Con M) {w : M} {x : M} {y : M} {z : M} :
c w xc y zc (w / y) (x / z)

Multiplicative congruence relations preserve division.

theorem AddCon.zsmul {M : Type u_1} [AddGroup M] (c : AddCon M) (n : ) {w : M} {x : M} :
c w xc (n w) (n x)

Additive congruence relations preserve integer scaling.

theorem Con.zpow {M : Type u_1} [Group M] (c : Con M) (n : ) {w : M} {x : M} :
c w xc (w ^ n) (x ^ n)

Multiplicative congruence relations preserve integer powers.

instance AddCon.hasNeg {M : Type u_1} [AddGroup M] (c : AddCon M) :
Neg c.Quotient

The negation induced on the quotient by an additive congruence relation on a type with a negation.

Equations
instance Con.hasInv {M : Type u_1} [Group M] (c : Con M) :
Inv c.Quotient

The inversion induced on the quotient by a congruence relation on a type with an inversion.

Equations
instance AddCon.hasSub {M : Type u_1} [AddGroup M] (c : AddCon M) :
Sub c.Quotient

The subtraction induced on the quotient by an additive congruence relation on a type with a subtraction.

Equations
instance Con.hasDiv {M : Type u_1} [Group M] (c : Con M) :
Div c.Quotient

The division induced on the quotient by a congruence relation on a type with a division.

Equations
instance AddCon.Quotient.zsmul {M : Type u_4} [AddGroup M] (c : AddCon M) :
SMul c.Quotient

The integer scaling induced on the quotient by a congruence relation on a type with a subtraction.

Equations
instance Con.zpowinst {M : Type u_1} [Group M] (c : Con M) :
Pow c.Quotient

The integer power induced on the quotient by a congruence relation on a type with a division.

Equations
  • c.zpowinst = { pow := fun (x : c.Quotient) (z : ) => Quotient.map' (fun (x : M) => x ^ z) x }
theorem AddCon.addGroup.proof_1 {M : Type u_1} [AddGroup M] (c : AddCon M) :
Function.Surjective Quotient.mk''
theorem AddCon.addGroup.proof_3 {M : Type u_1} [AddGroup M] (c : AddCon M) :
∀ (x x_1 : M), Quotient.mk'' (x + x_1) = Quotient.mk'' (x + x_1)
theorem AddCon.addGroup.proof_11 {M : Type u_1} [AddGroup M] (c : AddCon M) (n : ) (a : c.Quotient) :
theorem AddCon.addGroup.proof_4 {M : Type u_1} [AddGroup M] (c : AddCon M) :
∀ (x : M), Quotient.mk'' (-x) = Quotient.mk'' (-x)
theorem AddCon.addGroup.proof_8 {M : Type u_1} [AddGroup M] (c : AddCon M) (a : c.Quotient) (b : c.Quotient) :
a - b = a + -b
theorem AddCon.addGroup.proof_12 {M : Type u_1} [AddGroup M] (c : AddCon M) (a : c.Quotient) :
-a + a = 0
theorem AddCon.addGroup.proof_10 {M : Type u_1} [AddGroup M] (c : AddCon M) (n : ) (a : c.Quotient) :
theorem AddCon.addGroup.proof_6 {M : Type u_1} [AddGroup M] (c : AddCon M) :
∀ (x : M) (x_1 : ), Quotient.mk'' (x_1 x) = Quotient.mk'' (x_1 x)
theorem AddCon.addGroup.proof_7 {M : Type u_1} [AddGroup M] (c : AddCon M) :
∀ (x : M) (x_1 : ), Quotient.mk'' (x_1 x) = Quotient.mk'' (x_1 x)
theorem AddCon.addGroup.proof_5 {M : Type u_1} [AddGroup M] (c : AddCon M) :
∀ (x x_1 : M), Quotient.mk'' (x - x_1) = Quotient.mk'' (x - x_1)
instance AddCon.addGroup {M : Type u_1} [AddGroup M] (c : AddCon M) :
AddGroup c.Quotient

The quotient of an AddGroup by an additive congruence relation is an AddGroup.

Equations
theorem AddCon.addGroup.proof_9 {M : Type u_1} [AddGroup M] (c : AddCon M) (a : c.Quotient) :
instance Con.group {M : Type u_1} [Group M] (c : Con M) :
Group c.Quotient

The quotient of a group by a congruence relation is a group.

Equations
theorem AddCon.liftOnAddUnits.proof_2 {M : Type u_1} [AddMonoid M] {c : AddCon M} (x : M) (y : M) (hyx : y + x = 0) :
c (y + x) 0
def AddCon.liftOnAddUnits {M : Type u_1} {α : Type u_4} [AddMonoid M] {c : AddCon M} (u : AddUnits c.Quotient) (f : (x y : M) → c (x + y) 0c (y + x) 0α) (Hf : ∀ (x y : M) (hxy : c (x + y) 0) (hyx : c (y + x) 0) (x' y' : M) (hxy' : c (x' + y') 0) (hyx' : c (y' + x') 0), c x x'c y y'f x y hxy hyx = f x' y' hxy' hyx') :
α

In order to define a function (Con.Quotient c)ˣ → α on the units of Con.Quotient c, where c : Con M is a multiplicative congruence on a monoid, it suffices to define a function f that takes elements x y : M with proofs of c (x * y) 1 and c (y * x) 1, and returns an element of α provided that f x y _ _ = f x' y' _ _ whenever c x x' and c y y'.

Equations
theorem AddCon.liftOnAddUnits.proof_1 {M : Type u_1} [AddMonoid M] {c : AddCon M} (x : M) (y : M) (hxy : x + y = 0) :
c (x + y) 0
theorem AddCon.liftOnAddUnits.proof_5 {M : Type u_1} [AddMonoid M] {c : AddCon M} (u : AddUnits c.Quotient) :
u.neg + u = 0
theorem AddCon.liftOnAddUnits.proof_3 {M : Type u_1} {α : Type u_2} [AddMonoid M] {c : AddCon M} (f : (x y : M) → c (x + y) 0c (y + x) 0α) (Hf : ∀ (x y : M) (hxy : c (x + y) 0) (hyx : c (y + x) 0) (x' y' : M) (hxy' : c (x' + y') 0) (hyx' : c (y' + x') 0), c x x'c y y'f x y hxy hyx = f x' y' hxy' hyx') (x : M) (y : M) (x' : M) (y' : M) (hx : c x x') (hy : c y y') :
HEq ((fun (x y : M) (hxy : x + y = 0) (hyx : y + x = 0) => f x y ) x y) ((fun (x y : M) (hxy : x + y = 0) (hyx : y + x = 0) => f x y ) x' y')
theorem AddCon.liftOnAddUnits.proof_4 {M : Type u_1} [AddMonoid M] {c : AddCon M} (u : AddUnits c.Quotient) :
u + u.neg = 0
def Con.liftOnUnits {M : Type u_1} {α : Type u_4} [Monoid M] {c : Con M} (u : c.Quotientˣ) (f : (x y : M) → c (x * y) 1c (y * x) 1α) (Hf : ∀ (x y : M) (hxy : c (x * y) 1) (hyx : c (y * x) 1) (x' y' : M) (hxy' : c (x' * y') 1) (hyx' : c (y' * x') 1), c x x'c y y'f x y hxy hyx = f x' y' hxy' hyx') :
α

In order to define a function (Con.Quotient c)ˣ → α on the units of Con.Quotient c, where c : Con M is a multiplicative congruence on a monoid, it suffices to define a function f that takes elements x y : M with proofs of c (x * y) 1 and c (y * x) 1, and returns an element of α provided that f x y _ _ = f x' y' _ _ whenever c x x' and c y y'.

Equations
@[simp]
theorem AddCon.liftOnAddUnits_mk {M : Type u_1} {α : Type u_4} [AddMonoid M] {c : AddCon M} (f : (x y : M) → c (x + y) 0c (y + x) 0α) (Hf : ∀ (x y : M) (hxy : c (x + y) 0) (hyx : c (y + x) 0) (x' y' : M) (hxy' : c (x' + y') 0) (hyx' : c (y' + x') 0), c x x'c y y'f x y hxy hyx = f x' y' hxy' hyx') (x : M) (y : M) (hxy : x + y = 0) (hyx : y + x = 0) :
AddCon.liftOnAddUnits { val := x, neg := y, val_neg := hxy, neg_val := hyx } f Hf = f x y
@[simp]
theorem Con.liftOnUnits_mk {M : Type u_1} {α : Type u_4} [Monoid M] {c : Con M} (f : (x y : M) → c (x * y) 1c (y * x) 1α) (Hf : ∀ (x y : M) (hxy : c (x * y) 1) (hyx : c (y * x) 1) (x' y' : M) (hxy' : c (x' * y') 1) (hyx' : c (y' * x') 1), c x x'c y y'f x y hxy hyx = f x' y' hxy' hyx') (x : M) (y : M) (hxy : x * y = 1) (hyx : y * x = 1) :
Con.liftOnUnits { val := x, inv := y, val_inv := hxy, inv_val := hyx } f Hf = f x y
theorem AddCon.induction_on_addUnits {M : Type u_1} [AddMonoid M] {c : AddCon M} {p : AddUnits c.QuotientProp} (u : AddUnits c.Quotient) (H : ∀ (x y : M) (hxy : c (x + y) 0) (hyx : c (y + x) 0), p { val := x, neg := y, val_neg := , neg_val := }) :
p u
theorem Con.induction_on_units {M : Type u_1} [Monoid M] {c : Con M} {p : c.QuotientˣProp} (u : c.Quotientˣ) (H : ∀ (x y : M) (hxy : c (x * y) 1) (hyx : c (y * x) 1), p { val := x, inv := y, val_inv := , inv_val := }) :
p u
instance AddCon.instVAdd {α : Type u_4} {M : Type u_5} [AddZeroClass M] [VAdd α M] [VAddAssocClass α M M] (c : AddCon M) :
VAdd α c.Quotient
Equations
instance Con.instSMul {α : Type u_4} {M : Type u_5} [MulOneClass M] [SMul α M] [IsScalarTower α M M] (c : Con M) :
SMul α c.Quotient
Equations
theorem AddCon.coe_vadd {α : Type u_4} {M : Type u_5} [AddZeroClass M] [VAdd α M] [VAddAssocClass α M M] (c : AddCon M) (a : α) (x : M) :
(a +ᵥ x) = a +ᵥ x
theorem Con.coe_smul {α : Type u_4} {M : Type u_5} [MulOneClass M] [SMul α M] [IsScalarTower α M M] (c : Con M) (a : α) (x : M) :
(a x) = a x
theorem AddCon.addAction.proof_1 {α : Type u_2} {M : Type u_1} [AddMonoid α] [AddZeroClass M] [AddAction α M] [VAddAssocClass α M M] (c : AddCon M) (q : Quotient c.toSetoid) :
0 +ᵥ q = q
instance AddCon.addAction {α : Type u_4} {M : Type u_5} [AddMonoid α] [AddZeroClass M] [AddAction α M] [VAddAssocClass α M M] (c : AddCon M) :
AddAction α c.Quotient
Equations
theorem AddCon.addAction.proof_2 {α : Type u_2} {M : Type u_1} [AddMonoid α] [AddZeroClass M] [AddAction α M] [VAddAssocClass α M M] (c : AddCon M) :
∀ (x x_1 : α) (q : Quotient c.toSetoid), x + x_1 +ᵥ q = x +ᵥ (x_1 +ᵥ q)
instance Con.mulAction {α : Type u_4} {M : Type u_5} [Monoid α] [MulOneClass M] [MulAction α M] [IsScalarTower α M M] (c : Con M) :
MulAction α c.Quotient
Equations
instance Con.mulDistribMulAction {α : Type u_4} {M : Type u_5} [Monoid α] [Monoid M] [MulDistribMulAction α M] [IsScalarTower α M M] (c : Con M) :
MulDistribMulAction α c.Quotient
Equations