Further lemmas about units in a MonoidWithZero or a GroupWithZero. #
theorem
isLocalHom_of_exists_map_ne_one
{M : Type u_1}
{G₀ : Type u_3}
{F : Type u_6}
[Monoid M]
[GroupWithZero G₀]
[FunLike F G₀ M]
[MonoidHomClass F G₀ M]
{f : F}
(hf : ∃ (x : G₀), f x ≠ 1)
 :
instance
instIsLocalHomOfMonoidWithZeroHomClassOfNontrivial
{M₀ : Type u_2}
{G₀ : Type u_3}
{F : Type u_6}
[MonoidWithZero M₀]
[GroupWithZero G₀]
[FunLike F G₀ M₀]
[MonoidWithZeroHomClass F G₀ M₀]
[Nontrivial M₀]
(f : F)
 :
theorem
Commute.div_eq_div_iff
{G₀ : Type u_3}
[GroupWithZero G₀]
{a b c d : G₀}
(hbd : Commute b d)
(hb : b ≠ 0)
(hd : d ≠ 0)
 :
The MonoidWithZero version of div_eq_div_iff_mul_eq_mul.
theorem
Commute.mul_inv_eq_mul_inv_iff
{G₀ : Type u_3}
[GroupWithZero G₀]
{a b c d : G₀}
(hbd : Commute b d)
(hb : b ≠ 0)
(hd : d ≠ 0)
 :
The MonoidWithZero version of mul_inv_eq_mul_inv_iff_mul_eq_mul.
theorem
Commute.inv_mul_eq_inv_mul_iff
{G₀ : Type u_3}
[GroupWithZero G₀]
{a b c d : G₀}
(hbd : Commute b d)
(hb : b ≠ 0)
(hd : d ≠ 0)
 :
The MonoidWithZero version of inv_mul_eq_inv_mul_iff_mul_eq_mul.
theorem
map_ne_zero
{G₀ : Type u_3}
{M₀' : Type u_4}
{F : Type u_6}
[GroupWithZero G₀]
[MulZeroOneClass M₀']
[Nontrivial M₀']
[FunLike F G₀ M₀']
[MonoidWithZeroHomClass F G₀ M₀']
(f : F)
{a : G₀}
 :
@[simp]
theorem
map_eq_zero
{G₀ : Type u_3}
{M₀' : Type u_4}
{F : Type u_6}
[GroupWithZero G₀]
[MulZeroOneClass M₀']
[Nontrivial M₀']
[FunLike F G₀ M₀']
[MonoidWithZeroHomClass F G₀ M₀']
(f : F)
{a : G₀}
 :
theorem
eq_on_inv₀
{G₀ : Type u_3}
{M₀' : Type u_4}
{F' : Type u_7}
[GroupWithZero G₀]
[MonoidWithZero M₀']
[FunLike F' G₀ M₀']
{a : G₀}
[MonoidWithZeroHomClass F' G₀ M₀']
(f g : F')
(h : f a = g a)
 :
@[simp]
theorem
map_inv₀
{G₀ : Type u_3}
{G₀' : Type u_5}
{F : Type u_6}
[GroupWithZero G₀]
[GroupWithZero G₀']
[FunLike F G₀ G₀']
[MonoidWithZeroHomClass F G₀ G₀']
(f : F)
(a : G₀)
 :
A monoid homomorphism between groups with zeros sending 0 to 0 sends a⁻¹ to (f a)⁻¹.
@[simp]
theorem
map_div₀
{G₀ : Type u_3}
{G₀' : Type u_5}
{F : Type u_6}
[GroupWithZero G₀]
[GroupWithZero G₀']
[FunLike F G₀ G₀']
[MonoidWithZeroHomClass F G₀ G₀']
(f : F)
(a b : G₀)
 :
We define the inverse as a MonoidWithZeroHom by extending the inverse map by zero
on non-units.
Equations
- MonoidWithZero.inverse = { toFun := Ring.inverse, map_zero' := ⋯, map_one' := ⋯, map_mul' := ⋯ }
Instances For
@[simp]
@[simp]
Inversion on a commutative group with zero, considered as a monoid with zero homomorphism.
Equations
- invMonoidWithZeroHom = { toFun := (↑invMonoidHom).toFun, map_zero' := ⋯, map_one' := ⋯, map_mul' := ⋯ }
Instances For
@[simp]
theorem
map_zpow₀
{F : Type u_8}
{G₀ : Type u_9}
{G₀' : Type u_10}
[GroupWithZero G₀]
[GroupWithZero G₀']
[FunLike F G₀ G₀']
[MonoidWithZeroHomClass F G₀ G₀']
(f : F)
(x : G₀)
(n : ℤ)
 :
If a monoid homomorphism f between two GroupWithZeros maps 0 to 0, then it maps x^n,
n : ℤ, to (f x)^n.