Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit e83458c

Browse files
committed
feat(algebra/group_power): Add mul/add variants of powers_hom and friends (#4636)
1 parent c83c28a commit e83458c

File tree

1 file changed

+49
-0
lines changed

1 file changed

+49
-0
lines changed

src/algebra/group_power/lemmas.lean

Lines changed: 49 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@ import algebra.opposites
88
import data.list.basic
99
import data.int.cast
1010
import data.equiv.basic
11+
import data.equiv.mul_add
1112
import deprecated.group
1213

1314
/-!
@@ -459,6 +460,54 @@ by rw [← gmultiples_hom_symm_apply, ← gmultiples_hom_apply, equiv.apply_symm
459460

460461
/-! `add_monoid_hom.ext_int` is defined in `data.int.cast` -/
461462

463+
variables (M G A)
464+
465+
/-- If `M` is commutative, `powers_hom` is a multiplicative equivalence. -/
466+
def powers_mul_hom [comm_monoid M] : M ≃* (multiplicative ℕ →* M) :=
467+
{ map_mul' := λ a b, monoid_hom.ext $ by simp [mul_pow],
468+
..powers_hom M}
469+
470+
/-- If `M` is commutative, `gpowers_hom` is a multiplicative equivalence. -/
471+
def gpowers_mul_hom [comm_group G] : G ≃* (multiplicative ℤ →* G) :=
472+
{ map_mul' := λ a b, monoid_hom.ext $ by simp [mul_gpow],
473+
..gpowers_hom G}
474+
475+
/-- If `M` is commutative, `multiples_hom` is an additive equivalence. -/
476+
def multiples_add_hom [add_comm_monoid A] : A ≃+ (ℕ →+ A) :=
477+
{ map_add' := λ a b, add_monoid_hom.ext $ by simp [nsmul_add],
478+
..multiples_hom A}
479+
480+
/-- If `M` is commutative, `gmultiples_hom` is an additive equivalence. -/
481+
def gmultiples_add_hom [add_comm_group A] : A ≃+ (ℤ →+ A) :=
482+
{ map_add' := λ a b, add_monoid_hom.ext $ by simp [gsmul_add],
483+
..gmultiples_hom A}
484+
485+
variables {M G A}
486+
487+
@[simp] lemma powers_mul_hom_apply [comm_monoid M] (x : M) (n : multiplicative ℕ) :
488+
powers_mul_hom M x n = x ^ n.to_add := rfl
489+
490+
@[simp] lemma powers_mul_hom_symm_apply [comm_monoid M] (f : multiplicative ℕ →* M) :
491+
(powers_mul_hom M).symm f = f (multiplicative.of_add 1) := rfl
492+
493+
@[simp] lemma gpowers_mul_hom_apply [comm_group G] (x : G) (n : multiplicative ℤ) :
494+
gpowers_mul_hom G x n = x ^ n.to_add := rfl
495+
496+
@[simp] lemma gpowers_mul_hom_symm_apply [comm_group G] (f : multiplicative ℤ →* G) :
497+
(gpowers_mul_hom G).symm f = f (multiplicative.of_add 1) := rfl
498+
499+
@[simp] lemma multiples_add_hom_apply [add_comm_monoid A] (x : A) (n : ℕ) :
500+
multiples_add_hom A x n = n •ℕ x := rfl
501+
502+
@[simp] lemma multiples_add_hom_symm_apply [add_comm_monoid A] (f : ℕ →+ A) :
503+
(multiples_add_hom A).symm f = f 1 := rfl
504+
505+
@[simp] lemma gmultiples_add_hom_apply [add_comm_group A] (x : A) (n : ℤ) :
506+
gmultiples_add_hom A x n = n •ℤ x := rfl
507+
508+
@[simp] lemma gmultiples_add_hom_symm_apply [add_comm_group A] (f : ℤ →+ A) :
509+
(gmultiples_add_hom A).symm f = f 1 := rfl
510+
462511
/-!
463512
### Commutativity (again)
464513

0 commit comments

Comments
 (0)