Skip to content

Commit

Permalink
Add ext lemmas
Browse files Browse the repository at this point in the history
  • Loading branch information
eric-wieser committed Aug 29, 2021
1 parent d083a70 commit 5558fe8
Show file tree
Hide file tree
Showing 4 changed files with 29 additions and 8 deletions.
9 changes: 9 additions & 0 deletions src/algebra/direct_sum/algebra.lean
Expand Up @@ -131,6 +131,15 @@ def to_algebra
commutes' := λ r, (direct_sum.to_semiring_of _ _ _ _ _).trans (hcommutes r),
.. to_semiring (λ i, (f i).to_add_monoid_hom) hone @hmul}

/-- Two `alg_hom`s out of a direct sum are equal if they agree on the generators.
See note [partially-applied ext lemmas]. -/
@[ext]
lemma alg_hom_ext ⦃f g : (⨁ i, A i) →ₐ[R] B⦄
(h : ∀ i, f.to_linear_map.comp (lof _ _ A i) = g.to_linear_map.comp (lof _ _ A i)) : f = g :=
alg_hom.coe_ring_hom_injective $
direct_sum.ring_hom_ext $ λ i, add_monoid_hom.ext $ linear_map.congr_fun (h i)

end direct_sum

/-! ### Concrete instances -/
Expand Down
13 changes: 10 additions & 3 deletions src/algebra/direct_sum/module.lean
Expand Up @@ -56,6 +56,9 @@ dfinsupp.lmk
/-- Inclusion of each component into the direct sum. -/
def lof : Π i : ι, M i →ₗ[R] (⨁ i, M i) :=
dfinsupp.lsingle

lemma lof_eq_of (i : ι) (b : M i) : lof R ι M i b = of M i b := rfl

variables {ι M}

lemma single_eq_lof (i : ι) (b : M i) :
Expand Down Expand Up @@ -103,9 +106,13 @@ to_add_monoid.unique ψ.to_add_monoid_hom f

variables {ψ} {ψ' : (⨁ i, M i) →ₗ[R] N}

theorem to_module.ext (H : ∀ i, ψ.comp (lof R ι M i) = ψ'.comp (lof R ι M i)) (f : ⨁ i, M i) :
ψ f = ψ' f :=
by rw dfinsupp.lhom_ext' H
/-- Two `linear_map`s out of a direct sum are equal if they agree on the generators.
See note [partially-applied ext lemmas]. -/
@[ext]
theorem linear_map_ext ⦃ψ ψ' : (⨁ i, M i) →ₗ[R] N⦄
(H : ∀ i, ψ.comp (lof R ι M i) = ψ'.comp (lof R ι M i)) : ψ = ψ' :=
dfinsupp.lhom_ext' H

/--
The inclusion of a subset of the direct summands
Expand Down
9 changes: 9 additions & 0 deletions src/algebra/direct_sum/ring.lean
Expand Up @@ -634,6 +634,15 @@ def lift_ring_hom :
add_monoid_hom.comp_apply, to_semiring_coe_add_monoid_hom],
end}

/-- Two `ring_hom`s out of a direct sum are equal if they agree on the generators.
See note [partially-applied ext lemmas]. -/
@[ext]
lemma ring_hom_ext ⦃f g : (⨁ i, A i) →+* R⦄
(h : ∀ i, (↑f : (⨁ i, A i) →+ R).comp (of A i) = (↑g : (⨁ i, A i) →+ R).comp (of A i)) :
f = g :=
direct_sum.lift_ring_hom.symm.injective $ subtype.ext $ funext h

end to_semiring

end direct_sum
Expand Down
6 changes: 1 addition & 5 deletions src/linear_algebra/direct_sum/tensor_product.lean
Expand Up @@ -37,14 +37,10 @@ begin
(lift $ direct_sum.to_module R _ _ $ λ i₁, flip $ direct_sum.to_module R _ _ $ λ i₂,
flip $ curry $ direct_sum.lof R (ι₁ × ι₂) (λ i, M₁ i.1 ⊗[R] M₂ i.2) (i₁, i₂))
(direct_sum.to_module R _ _ $ λ i, map (direct_sum.lof R _ _ _) (direct_sum.lof R _ _ _))
(linear_map.ext $ direct_sum.to_module.ext _ $ λ i, mk_compr₂_inj $
linear_map.ext $ λ x₁, linear_map.ext $ λ x₂, _)
(mk_compr₂_inj $ linear_map.ext $ direct_sum.to_module.ext _ $ λ i₁, linear_map.ext $ λ x₁,
linear_map.ext $ direct_sum.to_module.ext _ $ λ i₂, linear_map.ext $ λ x₂, _),
_ _; [ext ⟨i₁, i₂⟩ x₁ x₂ : 4, ext i₁ i₂ x₁ x₂ : 5],
repeat { rw compr₂_apply <|> rw comp_apply <|> rw id_apply <|> rw mk_apply <|>
rw direct_sum.to_module_lof <|> rw map_tmul <|> rw lift.tmul <|> rw flip_apply <|>
rw curry_apply },
cases i; refl
end

@[simp] theorem direct_sum_lof_tmul_lof (i₁ : ι₁) (m₁ : M₁ i₁) (i₂ : ι₂) (m₂ : M₂ i₂) :
Expand Down

0 comments on commit 5558fe8

Please sign in to comment.