Skip to content

Commit 1639883

Browse files
urkuddagurtomas
authored andcommitted
chore(AddConstMap): don't extend FunLike (#14201)
Instead, take it as an argument as other bundled homs do. Also change the `npow` field in the `Monoid` instance so that `coe_pow` is a `rfl` lemma now.
1 parent c6b934b commit 1639883

File tree

1 file changed

+19
-15
lines changed

1 file changed

+19
-15
lines changed

Mathlib/Algebra/AddConstMap/Basic.lean

+19-15
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@ import Mathlib.Algebra.Module.Defs
88
import Mathlib.Algebra.Order.Archimedean
99
import Mathlib.Algebra.Order.Group.Instances
1010
import Mathlib.GroupTheory.GroupAction.Pi
11+
import Mathlib.Logic.Function.Iterate
1112

1213
/-!
1314
# Maps (semi)conjugating a shift to a shift
@@ -48,7 +49,7 @@ Note that `a` and `b` are `outParam`s,
4849
so one should not add instances like
4950
`[AddConstMapClass F G H a b] : AddConstMapClass F G H (-a) (-b)`. -/
5051
class AddConstMapClass (F : Type*) (G H : outParam Type*) [Add G] [Add H]
51-
(a : outParam G) (b : outParam H) extends DFunLike F G fun _ ↦ H where
52+
(a : outParam G) (b : outParam H) [FunLike F G H] : Prop where
5253
/-- A map of `AddConstMapClass` class semiconjugates shift by `a` to the shift by `b`:
5354
`∀ x, f (x + a) = f x + b`. -/
5455
map_add_const (f : F) (x : G) : f (x + a) = f x + b
@@ -63,7 +64,7 @@ In this section we prove properties like `f (x + n • a) = f x + n • b`.
6364

6465
attribute [simp] map_add_const
6566

66-
variable {F G H : Type*} {a : G} {b : H}
67+
variable {F G H : Type*} [FunLike F G H] {a : G} {b : H}
6768

6869
protected theorem semiconj [Add G] [Add H] [AddConstMapClass F G H a b] (f : F) :
6970
Semiconj f (· + a) (· + b) :=
@@ -227,7 +228,7 @@ theorem map_int_add [AddCommGroupWithOne G] [AddGroupWithOne H] [AddConstMapClas
227228
(f : F) (n : ℤ) (x : G) : f (↑n + x) = f x + n := by simp
228229

229230
theorem map_fract {R : Type*} [LinearOrderedRing R] [FloorRing R] [AddGroup H]
230-
[AddConstMapClass F R H 1 b] (f : F) (x : R) :
231+
[FunLike F R H] [AddConstMapClass F R H 1 b] (f : F) (x : R) :
231232
f (Int.fract x) = f x - ⌊x⌋ • b :=
232233
map_sub_int' ..
233234

@@ -309,13 +310,16 @@ variable {G H : Type*} [Add G] [Add H] {a : G} {b : H}
309310
### Coercion to function
310311
-/
311312

312-
instance : AddConstMapClass (G →+c[a, b] H) G H a b where
313+
instance : FunLike (G →+c[a, b] H) G H where
313314
coe := AddConstMap.toFun
314315
coe_injective' | ⟨_, _⟩, ⟨_, _⟩, rfl => rfl
315-
map_add_const f := f.map_add_const'
316316

317317
@[simp] theorem coe_mk (f : G → H) (hf) : ⇑(mk f hf : G →+c[a, b] H) = f := rfl
318318
@[simp] theorem mk_coe (f : G →+c[a, b] H) : mk f f.2 = f := rfl
319+
@[simp] theorem toFun_eq_coe (f : G →+c[a, b] H) : f.toFun = f := rfl
320+
321+
instance : AddConstMapClass (G →+c[a, b] H) G H a b where
322+
map_add_const f := f.map_add_const'
319323

320324
@[ext] protected theorem ext {f g : G →+c[a, b] H} (h : ∀ x, f x = g x) : f = g :=
321325
DFunLike.ext _ _ h
@@ -369,24 +373,24 @@ instance {K : Type*} [AddMonoid K] [AddAction K H] [VAddAssocClass K H H] :
369373
### Monoid structure on endomorphisms `G →+c[a, a] G`
370374
-/
371375

372-
instance : Monoid (G →+c[a, a] G) where
373-
mul := comp
374-
one := .id
375-
mul_assoc _ _ _ := rfl
376-
one_mul := id_comp
377-
mul_one := comp_id
376+
instance : Mul (G →+c[a, a] G) := ⟨comp⟩
377+
instance : One (G →+c[a, a] G) := ⟨.id⟩
378+
379+
instance : Pow (G →+c[a, a] G) ℕ where
380+
pow f n := ⟨f^[n], Commute.iterate_left (AddConstMapClass.semiconj f) _⟩
381+
382+
instance : Monoid (G →+c[a, a] G) :=
383+
DFunLike.coe_injective.monoid (M₂ := Function.End G) _ rfl (fun _ _ ↦ rfl) fun _ _ ↦ rfl
378384

379385
theorem mul_def (f g : G →+c[a, a] G) : f * g = f.comp g := rfl
380386
@[simp] theorem coe_mul (f g : G →+c[a, a] G) : ⇑(f * g) = f ∘ g := rfl
381387

382388
theorem one_def : (1 : G →+c[a, a] G) = .id := rfl
383389
@[simp] theorem coe_one : ⇑(1 : G →+c[a, a] G) = id := rfl
384390

385-
theorem coe_pow (f : G →+c[a, a] G) (n : ℕ) : ⇑(f ^ n) = f^[n] :=
386-
hom_coe_pow _ rfl (fun _ _ ↦ rfl) _ _
391+
@[simp] theorem coe_pow (f : G →+c[a, a] G) (n : ℕ) : ⇑(f ^ n) = f^[n] := rfl
387392

388-
theorem pow_apply (f : G →+c[a, a] G) (n : ℕ) (x : G) : (f ^ n) x = f^[n] x :=
389-
congr_fun (coe_pow f n) x
393+
theorem pow_apply (f : G →+c[a, a] G) (n : ℕ) (x : G) : (f ^ n) x = f^[n] x := rfl
390394

391395
end Add
392396

0 commit comments

Comments
 (0)