Skip to content

Commit

Permalink
feat(linear_algebra): composing with a linear equivalence does not ch…
Browse files Browse the repository at this point in the history
…ange the image (#6816)

I also did some minor reorganisation in order to relax some typeclass arguments.
  • Loading branch information
TwoFX committed Mar 22, 2021
1 parent e54f633 commit ffca31a
Showing 1 changed file with 35 additions and 8 deletions.
43 changes: 35 additions & 8 deletions src/linear_algebra/basic.lean
Expand Up @@ -1701,15 +1701,11 @@ end ring
end submodule

namespace linear_map
variables [ring R] [add_comm_group M] [add_comm_group M₂] [add_comm_group M₃]
variables [module R M] [module R M₂] [module R M₃]

lemma range_mkq_comp (f : M →ₗ[R] M₂) : f.range.mkq.comp f = 0 :=
linear_map.ext $ λ x, by simp
section semiring

lemma ker_le_range_iff {f : M →ₗ[R] M₂} {g : M₂ →ₗ[R] M₃} :
g.ker ≤ f.range ↔ f.range.mkq.comp g.ker.subtype = 0 :=
by rw [←range_le_ker_iff, submodule.ker_mkq, submodule.range_subtype]
variables [semiring R] [add_comm_monoid M] [add_comm_monoid M₂] [add_comm_monoid M₃]
variables [semimodule R M] [semimodule R M₂] [semimodule R M₃]

/-- A monomorphism is injective. -/
lemma ker_eq_bot_of_cancel {f : M →ₗ[R] M₂}
Expand All @@ -1720,6 +1716,28 @@ begin
exact range_zero
end

lemma range_comp_of_range_eq_top {f : M →ₗ[R] M₂} (g : M₂ →ₗ[R] M₃) (hf : range f = ⊤) :
range (g.comp f) = range g :=
by rw [range_comp, hf, submodule.map_top]

lemma ker_comp_of_ker_eq_bot (f : M →ₗ[R] M₂) {g : M₂ →ₗ[R] M₃} (hg : ker g = ⊥) :
ker (g.comp f) = ker f :=
by rw [ker_comp, hg, submodule.comap_bot]

end semiring

section ring

variables [ring R] [add_comm_monoid M] [add_comm_group M₂] [add_comm_monoid M₃]
variables [semimodule R M] [semimodule R M₂] [semimodule R M₃]

lemma range_mkq_comp (f : M →ₗ[R] M₂) : f.range.mkq.comp f = 0 :=
linear_map.ext $ λ x, by simp

lemma ker_le_range_iff {f : M →ₗ[R] M₂} {g : M₂ →ₗ[R] M₃} :
g.ker ≤ f.range ↔ f.range.mkq.comp g.ker.subtype = 0 :=
by rw [←range_le_ker_iff, submodule.ker_mkq, submodule.range_subtype]

/-- An epimorphism is surjective. -/
lemma range_eq_top_of_cancel {f : M →ₗ[R] M₂}
(h : ∀ (u v : M₂ →ₗ[R] f.range.quotient), u.comp f = v.comp f → u = v) : f.range = ⊤ :=
Expand All @@ -1729,6 +1747,8 @@ begin
exact ker_zero
end

end ring

end linear_map

@[simp] lemma linear_map.range_range_restrict [semiring R] [add_comm_monoid M] [add_comm_monoid M₂]
Expand Down Expand Up @@ -1789,7 +1809,8 @@ end uncurry

section
variables {semimodule_M : semimodule R M} {semimodule_M₂ : semimodule R M₂}
variables (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M) (e : M ≃ₗ[R] M₂)
{semimodule_M₃ : semimodule R M₃}
variables (f : M →ₗ[R] M₂) (g : M₂ →ₗ[R] M) (e : M ≃ₗ[R] M₂) (h : M₂ →ₗ[R] M₃) (l : M₃ →ₗ[R] M)

variables (p q : submodule R M)

Expand Down Expand Up @@ -1853,6 +1874,12 @@ end
@[simp] protected theorem ker : (e : M →ₗ[R] M₂).ker = ⊥ :=
linear_map.ker_eq_bot_of_injective e.to_equiv.injective

@[simp] theorem range_comp : (h.comp (e : M →ₗ[R] M₂)).range = h.range :=
linear_map.range_comp_of_range_eq_top _ e.range

@[simp] theorem ker_comp : ((e : M →ₗ[R] M₂).comp l).ker = l.ker :=
linear_map.ker_comp_of_ker_eq_bot _ e.ker

variables {f g}

/-- An linear map `f : M →ₗ[R] M₂` with a left-inverse `g : M₂ →ₗ[R] M` defines a linear equivalence
Expand Down

0 comments on commit ffca31a

Please sign in to comment.