Lifting groups with zero along injective/surjective maps #
Pull back a MulZeroClass
instance along an injective function.
See note [reducible non-instances].
Equations
- Function.Injective.mulZeroClass f hf zero mul = { mul := fun (x1 x2 : M₀') => x1 * x2, zero := 0, zero_mul := ⋯, mul_zero := ⋯ }
Instances For
Push forward a MulZeroClass
instance along a surjective function.
See note [reducible non-instances].
Equations
- Function.Surjective.mulZeroClass f hf zero mul = { mul := fun (x1 x2 : M₀') => x1 * x2, zero := 0, zero_mul := ⋯, mul_zero := ⋯ }
Instances For
Pull back a NoZeroDivisors
instance along an injective function.
Pull back a MulZeroOneClass
instance along an injective function.
See note [reducible non-instances].
Equations
- One or more equations did not get rendered due to their size.
Instances For
Push forward a MulZeroOneClass
instance along a surjective function.
See note [reducible non-instances].
Equations
- One or more equations did not get rendered due to their size.
Instances For
Pull back a SemigroupWithZero
along an injective function.
See note [reducible non-instances].
Equations
- Function.Injective.semigroupWithZero f hf zero mul = { toMul := MulZeroClass.toMul, mul_assoc := ⋯, toZero := MulZeroClass.toZero, zero_mul := ⋯, mul_zero := ⋯ }
Instances For
Push forward a SemigroupWithZero
along a surjective function.
See note [reducible non-instances].
Equations
- Function.Surjective.semigroupWithZero f hf zero mul = { toMul := MulZeroClass.toMul, mul_assoc := ⋯, toZero := MulZeroClass.toZero, zero_mul := ⋯, mul_zero := ⋯ }
Instances For
Pull back a MonoidWithZero
along an injective function.
See note [reducible non-instances].
Equations
- Function.Injective.monoidWithZero f hf zero one mul npow = { toMonoid := Function.Injective.monoid f hf one mul npow, toZero := MulZeroClass.toZero, zero_mul := ⋯, mul_zero := ⋯ }
Instances For
Push forward a MonoidWithZero
along a surjective function.
See note [reducible non-instances].
Equations
- Function.Surjective.monoidWithZero f hf zero one mul npow = { toMonoid := Function.Surjective.monoid f hf one mul npow, toZero := MulZeroClass.toZero, zero_mul := ⋯, mul_zero := ⋯ }
Instances For
Pull back a CommMonoidWithZero
along an injective function.
See note [reducible non-instances].
Equations
- Function.Injective.commMonoidWithZero f hf zero one mul npow = { toCommMonoid := Function.Injective.commMonoid f hf one mul npow, toZero := MulZeroClass.toZero, zero_mul := ⋯, mul_zero := ⋯ }
Instances For
Push forward a CommMonoidWithZero
along a surjective function.
See note [reducible non-instances].
Equations
- Function.Surjective.commMonoidWithZero f hf zero one mul npow = { toCommMonoid := Function.Surjective.commMonoid f hf one mul npow, toZero := MulZeroClass.toZero, zero_mul := ⋯, mul_zero := ⋯ }
Instances For
Pull back a CancelMonoidWithZero
along an injective function.
See note [reducible non-instances].
Equations
- One or more equations did not get rendered due to their size.
Instances For
Pull back a CancelCommMonoidWithZero
along an injective function.
See note [reducible non-instances].
Equations
- Function.Injective.cancelCommMonoidWithZero f hf zero one mul npow = { toCommMonoidWithZero := Function.Injective.commMonoidWithZero f hf zero one mul npow, toIsLeftCancelMulZero := ⋯ }
Instances For
Pull back a GroupWithZero
along an injective function.
See note [reducible non-instances].
Equations
- One or more equations did not get rendered due to their size.
Instances For
Push forward a GroupWithZero
along a surjective function.
See note [reducible non-instances].
Equations
- One or more equations did not get rendered due to their size.
Instances For
Pull back a CommGroupWithZero
along an injective function.
See note [reducible non-instances].
Equations
- One or more equations did not get rendered due to their size.
Instances For
Push forward a CommGroupWithZero
along a surjective function.
See note [reducible non-instances].
Equations
- One or more equations did not get rendered due to their size.