Skip to content

Commit

Permalink
feat(algebra/module/linear_map): Add linear_map.iterate (#6377)
Browse files Browse the repository at this point in the history
Co-authored-by: Alex J. Best <alex.j.best@gmail.com>





Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
  • Loading branch information
Ruben-VandeVelde and semorrison committed Mar 21, 2021
1 parent e20c730 commit 4cc4207
Show file tree
Hide file tree
Showing 2 changed files with 32 additions and 6 deletions.
6 changes: 6 additions & 0 deletions src/algebra/module/linear_map.lean
Expand Up @@ -217,6 +217,12 @@ def comp : M →ₗ[R] M₃ := ⟨f ∘ g, by simp, by simp⟩
@[norm_cast]
lemma comp_coe : (f : M₂ → M₃) ∘ (g : M → M₂) = f.comp g := rfl

@[simp] theorem comp_id : f.comp id = f :=
linear_map.ext $ λ x, rfl

@[simp] theorem id_comp : id.comp f = f :=
linear_map.ext $ λ x, rfl

end

/-- If a function `g` is a left and right inverse of a linear map `f`, then `g` is linear itself. -/
Expand Down
32 changes: 26 additions & 6 deletions src/linear_algebra/basic.lean
Expand Up @@ -106,12 +106,6 @@ variables [semimodule R M] [semimodule R M₂] [semimodule R M₃] [semimodule R
variables (f g : M →ₗ[R] M₂)
include R

@[simp] theorem comp_id : f.comp id = f :=
linear_map.ext $ λ x, rfl

@[simp] theorem id_comp : id.comp f = f :=
linear_map.ext $ λ x, rfl

theorem comp_assoc (g : M₂ →ₗ[R] M₃) (h : M₃ →ₗ[R] M₄) : (h.comp g).comp f = h.comp (g.comp f) :=
rfl

Expand Down Expand Up @@ -250,6 +244,32 @@ end
lemma coe_pow (f : M →ₗ[R] M) (n : ℕ) : ⇑(f^n) = (f^[n]) :=
by { ext m, apply pow_apply, }

section
variables {f' : M →ₗ[R] M}

lemma iterate_succ (n : ℕ) : (f' ^ (n + 1)) = comp (f' ^ n) f' :=
by rw [pow_succ', mul_eq_comp]

lemma iterate_surjective (h : surjective f') : ∀ n : ℕ, surjective ⇑(f' ^ n)
| 0 := surjective_id
| (n + 1) := by { rw [iterate_succ], exact surjective.comp (iterate_surjective n) h, }

lemma iterate_injective (h : injective f') : ∀ n : ℕ, injective ⇑(f' ^ n)
| 0 := injective_id
| (n + 1) := by { rw [iterate_succ], exact injective.comp (iterate_injective n) h, }

lemma iterate_bijective (h : bijective f') : ∀ n : ℕ, bijective ⇑(f' ^ n)
| 0 := bijective_id
| (n + 1) := by { rw [iterate_succ], exact bijective.comp (iterate_bijective n) h, }

lemma injective_of_iterate_injective {n : ℕ} (hn : n ≠ 0) (h : injective ⇑(f' ^ n)) :
injective f' :=
begin
rw [← nat.succ_pred_eq_of_pos (pos_iff_ne_zero.mpr hn), iterate_succ, ←comp_coe] at h,
exact injective.of_comp h,
end
end

section
open_locale classical

Expand Down

0 comments on commit 4cc4207

Please sign in to comment.