Documentation

Mathlib.Algebra.Hom.Iterate

Iterates of monoid and ring homomorphisms #

Iterate of a monoid/ring homomorphism is a monoid/ring homomorphism but it has a wrong type, so Lean can't apply lemmas like MonoidHom.map_one to f^[n] 1. Though it is possible to define a monoid structure on the endomorphisms, quite often we do not want to convert from M →* M to Monoid.End M and from f^[n] to f^n just to apply a simple lemma.

So, we restate standard *Hom.map_* lemmas under names *Hom.iterate_map_*.

We also prove formulas for iterates of add/mul left/right.

Tags #

homomorphism, iterate

theorem hom_coe_pow {M : Type u_1} {F : Type u_5} [Monoid F] (c : FMM) (h1 : c 1 = id) (hmul : ∀ (f g : F), c (f * g) = c f c g) (f : F) (n : ) :
c (f ^ n) = (c f)^[n]

An auxiliary lemma that can be used to prove ⇑(f ^ n) = ⇑f^[n].

@[simp]
theorem iterate_map_add {M : Type u_5} {F : Type u_6} [AddZeroClass M] (f : F) (n : ) (x : M) (y : M) [AddHomClass F M M] :
@[simp]
theorem iterate_map_mul {M : Type u_5} {F : Type u_6} [MulOneClass M] (f : F) (n : ) (x : M) (y : M) [MulHomClass F M M] :
@[simp]
theorem AddMonoidHom.iterate_map_zero {M : Type u_1} [AddZeroClass M] (f : M →+ M) (n : ) :
@[simp]
theorem MonoidHom.iterate_map_one {M : Type u_1} [MulOneClass M] (f : M →* M) (n : ) :
@[simp]
theorem AddMonoidHom.iterate_map_neg {G : Type u_3} [AddGroup G] (f : G →+ G) (n : ) (x : G) :
@[simp]
theorem MonoidHom.iterate_map_inv {G : Type u_3} [Group G] (f : G →* G) (n : ) (x : G) :
@[simp]
theorem AddMonoidHom.iterate_map_sub {G : Type u_3} [AddGroup G] (f : G →+ G) (n : ) (x : G) (y : G) :
@[simp]
theorem MonoidHom.iterate_map_div {G : Type u_3} [Group G] (f : G →* G) (n : ) (x : G) (y : G) :
theorem MonoidHom.iterate_map_pow {M : Type u_1} [Monoid M] (f : M →* M) (n : ) (a : M) (m : ) :
(f)^[n] (a ^ m) = (f)^[n] a ^ m
theorem MonoidHom.iterate_map_zpow {G : Type u_3} [Group G] (f : G →* G) (n : ) (a : G) (m : ) :
(f)^[n] (a ^ m) = (f)^[n] a ^ m
theorem MonoidHom.coe_pow {M : Type u_5} [CommMonoid M] (f : Monoid.End M) (n : ) :
↑(f ^ n) = (f)^[n]
theorem Monoid.End.coe_pow {M : Type u_5} [Monoid M] (f : Monoid.End M) (n : ) :
↑(f ^ n) = (f)^[n]
theorem AddMonoidHom.iterate_map_smul {M : Type u_1} [AddMonoid M] (f : M →+ M) (n : ) (m : ) (x : M) :
theorem AddMonoidHom.iterate_map_nsmul {M : Type u_1} [AddMonoid M] (f : M →+ M) (n : ) (m : ) (a : M) :
theorem AddMonoidHom.iterate_map_zsmul {G : Type u_3} [AddGroup G] (f : G →+ G) (n : ) (m : ) (x : G) :
theorem AddMonoid.End.coe_pow {A : Type u_5} [AddMonoid A] (f : AddMonoid.End A) (n : ) :
↑(f ^ n) = (f)^[n]
theorem RingHom.coe_pow {R : Type u_5} [Semiring R] (f : R →+* R) (n : ) :
↑(f ^ n) = (f)^[n]
theorem RingHom.iterate_map_one {R : Type u_5} [Semiring R] (f : R →+* R) (n : ) :
theorem RingHom.iterate_map_zero {R : Type u_5} [Semiring R] (f : R →+* R) (n : ) :
theorem RingHom.iterate_map_pow {R : Type u_5} [Semiring R] (f : R →+* R) (a : R) (n : ) (m : ) :
(f)^[n] (a ^ m) = (f)^[n] a ^ m
theorem RingHom.iterate_map_smul {R : Type u_5} [Semiring R] (f : R →+* R) (n : ) (m : ) (x : R) :
theorem RingHom.iterate_map_sub {R : Type u_5} [Ring R] (f : R →+* R) (n : ) (x : R) (y : R) :
theorem RingHom.iterate_map_neg {R : Type u_5} [Ring R] (f : R →+* R) (n : ) (x : R) :
theorem RingHom.iterate_map_zsmul {R : Type u_5} [Ring R] (f : R →+* R) (n : ) (m : ) (x : R) :
@[simp]
theorem vadd_iterate {G : Type u_3} {H : Type u_4} [AddMonoid G] (a : G) (n : ) [AddAction G H] :
(fun x => a +ᵥ x)^[n] = fun x => n a +ᵥ x
@[simp]
theorem smul_iterate {G : Type u_3} {H : Type u_4} [Monoid G] (a : G) (n : ) [MulAction G H] :
(fun x => a x)^[n] = fun x => a ^ n x
@[simp]
theorem add_left_iterate {G : Type u_3} [AddMonoid G] (a : G) (n : ) :
(fun x => a + x)^[n] = fun x => n a + x
@[simp]
theorem mul_left_iterate {G : Type u_3} [Monoid G] (a : G) (n : ) :
(fun x => a * x)^[n] = fun x => a ^ n * x
@[simp]
theorem add_right_iterate {G : Type u_3} [AddMonoid G] (a : G) (n : ) :
(fun x => x + a)^[n] = fun x => x + n a
@[simp]
theorem mul_right_iterate {G : Type u_3} [Monoid G] (a : G) (n : ) :
(fun x => x * a)^[n] = fun x => x * a ^ n
theorem add_right_iterate_apply_zero {G : Type u_3} [AddMonoid G] (a : G) (n : ) :
theorem mul_right_iterate_apply_one {G : Type u_3} [Monoid G] (a : G) (n : ) :
(fun x => x * a)^[n] 1 = a ^ n
@[simp]
theorem nsmul_iterate {G : Type u_3} [AddMonoid G] (n : ) (j : ) :
(fun x => n x)^[j] = fun x => n ^ j x
@[simp]
theorem pow_iterate {G : Type u_3} [Monoid G] (n : ) (j : ) :
(fun x => x ^ n)^[j] = fun x => x ^ n ^ j
@[simp]
theorem zsmul_iterate {G : Type u_3} [AddGroup G] (n : ) (j : ) :
(fun x => n x)^[j] = fun x => n ^ j x
@[simp]
theorem zpow_iterate {G : Type u_3} [Group G] (n : ) (j : ) :
(fun x => x ^ n)^[j] = fun x => x ^ n ^ j
theorem AddSemiconjBy.function_semiconj_add_left {G : Type u_3} [AddSemigroup G] {a : G} {b : G} {c : G} (h : AddSemiconjBy a b c) :
Function.Semiconj (fun x => a + x) (fun x => b + x) fun x => c + x
theorem SemiconjBy.function_semiconj_mul_left {G : Type u_3} [Semigroup G] {a : G} {b : G} {c : G} (h : SemiconjBy a b c) :
Function.Semiconj (fun x => a * x) (fun x => b * x) fun x => c * x
theorem AddCommute.function_commute_add_left {G : Type u_3} [AddSemigroup G] {a : G} {b : G} (h : AddCommute a b) :
Function.Commute (fun x => a + x) fun x => b + x
theorem Commute.function_commute_mul_left {G : Type u_3} [Semigroup G] {a : G} {b : G} (h : Commute a b) :
Function.Commute (fun x => a * x) fun x => b * x
theorem AddSemiconjBy.function_semiconj_add_right_swap {G : Type u_3} [AddSemigroup G] {a : G} {b : G} {c : G} (h : AddSemiconjBy a b c) :
Function.Semiconj (fun x => x + a) (fun x => x + c) fun x => x + b
theorem SemiconjBy.function_semiconj_mul_right_swap {G : Type u_3} [Semigroup G] {a : G} {b : G} {c : G} (h : SemiconjBy a b c) :
Function.Semiconj (fun x => x * a) (fun x => x * c) fun x => x * b
theorem AddCommute.function_commute_add_right {G : Type u_3} [AddSemigroup G] {a : G} {b : G} (h : AddCommute a b) :
Function.Commute (fun x => x + a) fun x => x + b
theorem Commute.function_commute_mul_right {G : Type u_3} [Semigroup G] {a : G} {b : G} (h : Commute a b) :
Function.Commute (fun x => x * a) fun x => x * b