The subgroup generated by an element.
Equations
- Subgroup.zpowers g = { carrier := Set.range fun (x : ℤ) => g ^ x, mul_mem' := ⋯, one_mem' := ⋯, inv_mem' := ⋯ }
Instances For
The additive subgroup generated by an element.
Equations
- AddSubgroup.zmultiples g = { carrier := Set.range fun (x : ℤ) => x • g, add_mem' := ⋯, zero_mem' := ⋯, neg_mem' := ⋯ }
Instances For
@[simp]
theorem
AddSubgroup.coe_zmultiples
{G : Type u_1}
[AddGroup G]
(g : G)
:
↑(zmultiples g) = Set.range fun (x : ℤ) => x • g
noncomputable instance
Subgroup.decidableMemZPowers
{G : Type u_1}
[Group G]
{a : G}
:
DecidablePred fun (x : G) => x ∈ zpowers a
Equations
- Subgroup.decidableMemZPowers = Classical.decPred fun (x : G) => x ∈ Subgroup.zpowers a
noncomputable instance
AddSubgroup.decidableMemZMultiples
{G : Type u_1}
[AddGroup G]
{a : G}
:
DecidablePred fun (x : G) => x ∈ zmultiples a
Equations
- AddSubgroup.decidableMemZMultiples = Classical.decPred fun (x : G) => x ∈ AddSubgroup.zmultiples a
theorem
AddSubgroup.zmultiples_eq_closure
{G : Type u_1}
[AddGroup G]
(g : G)
:
zmultiples g = closure {g}
@[simp]
theorem
AddSubgroup.zsmul_mem_zmultiples
{G : Type u_1}
[AddGroup G]
(g : G)
(k : ℤ)
:
k • g ∈ zmultiples g
@[simp]
theorem
AddSubgroup.nsmul_mem_zmultiples
{G : Type u_1}
[AddGroup G]
(g : G)
(k : ℕ)
:
k • g ∈ zmultiples g
@[simp]
theorem
AddSubgroup.forall_zmultiples
{G : Type u_1}
[AddGroup G]
{x : G}
{p : ↥(zmultiples x) → Prop}
:
(∀ (g : ↥(zmultiples x)), p g) ↔ ∀ (m : ℤ), p ⟨m • x, ⋯⟩
@[simp]
theorem
AddSubgroup.exists_zmultiples
{G : Type u_1}
[AddGroup G]
{x : G}
{p : ↥(zmultiples x) → Prop}
:
(∃ (g : ↥(zmultiples x)), p g) ↔ ∃ (m : ℤ), p ⟨m • x, ⋯⟩
theorem
AddSubgroup.forall_mem_zmultiples
{G : Type u_1}
[AddGroup G]
{x : G}
{p : G → Prop}
:
(∀ g ∈ zmultiples x, p g) ↔ ∀ (m : ℤ), p (m • x)
theorem
AddSubgroup.exists_mem_zmultiples
{G : Type u_1}
[AddGroup G]
{x : G}
{p : G → Prop}
:
(∃ g ∈ zmultiples x, p g) ↔ ∃ (m : ℤ), p (m • x)
@[simp]
theorem
MonoidHom.map_zpowers
{G : Type u_1}
[Group G]
{N : Type u_3}
[Group N]
(f : G →* N)
(x : G)
:
Subgroup.map f (Subgroup.zpowers x) = Subgroup.zpowers (f x)
@[simp]
theorem
AddMonoidHom.map_zmultiples
{G : Type u_1}
[AddGroup G]
{N : Type u_3}
[AddGroup N]
(f : G →+ N)
(x : G)
:
AddSubgroup.map f (AddSubgroup.zmultiples x) = AddSubgroup.zmultiples (f x)
theorem
ofMul_image_zpowers_eq_zmultiples_ofMul
{G : Type u_1}
[Group G]
{x : G}
:
⇑Additive.ofMul '' ↑(Subgroup.zpowers x) = ↑(AddSubgroup.zmultiples (Additive.ofMul x))
theorem
ofAdd_image_zmultiples_eq_zpowers_ofAdd
{A : Type u_2}
[AddGroup A]
{x : A}
:
⇑Multiplicative.ofAdd '' ↑(AddSubgroup.zmultiples x) = ↑(Subgroup.zpowers (Multiplicative.ofAdd x))
instance
AddSubgroup.zmultiples_isCommutative
{G : Type u_1}
[AddGroup G]
(g : G)
:
(zmultiples g).IsCommutative
@[simp]
theorem
AddSubgroup.zmultiples_le
{G : Type u_1}
[AddGroup G]
{g : G}
{H : AddSubgroup G}
:
zmultiples g ≤ H ↔ g ∈ H
Alias of the reverse direction of Subgroup.zpowers_le
.
theorem
AddSubgroup.zmultiples_le_of_mem
{G : Type u_1}
[AddGroup G]
{g : G}
{H : AddSubgroup G}
:
g ∈ H → zmultiples g ≤ H
Alias of the reverse direction of AddSubgroup.zmultiples_le
.
@[simp]
@[simp]
@[simp]
theorem
AddSubgroup.zmultiples_neg
{G : Type u_1}
[AddGroup G]
{g : G}
:
zmultiples (-g) = zmultiples g