Skip to content

Commit

Permalink
fix(algebra/direct_sum): change ring_hom_ext to not duplicate `ring…
Browse files Browse the repository at this point in the history
…_hom_ext'` (#10640)

These two lemmas differed only in the explicitness of their binders.
The statement of the unprimed version has been changed to be fully applied.

This also renames `alg_hom_ext` to `alg_hom_ext'` to make way for the fully applied version. This is consistent with `direct_sum.add_hom_ext` vs `direct_sum.add_hom_ext'`.
  • Loading branch information
eric-wieser committed Dec 8, 2021
1 parent b495fdf commit e5ba338
Show file tree
Hide file tree
Showing 2 changed files with 12 additions and 14 deletions.
8 changes: 5 additions & 3 deletions src/algebra/direct_sum/algebra.lean
Expand Up @@ -110,10 +110,12 @@ def to_algebra
See note [partially-applied ext lemmas]. -/
@[ext]
lemma alg_hom_ext ⦃f g : (⨁ i, A i) →ₐ[R] B⦄
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)
alg_hom.to_linear_map_injective $ direct_sum.linear_map_ext _ h

lemma alg_hom_ext ⦃f g : (⨁ i, A i) →ₐ[R] B⦄ (h : ∀ i x, f (of A i x) = g (of A i x)) : f = g :=
alg_hom_ext' R A $ λ i, linear_map.ext $ h i

end direct_sum

Expand Down
18 changes: 7 additions & 11 deletions src/algebra/direct_sum/ring.lean
Expand Up @@ -404,10 +404,15 @@ then they are equal.
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 _ i) = (G : (⨁ i, A i) →+ R).comp (of _ i)) : F = G :=
lemma ring_hom_ext' F G : (⨁ i, A i) →+* R
(h : ∀ i, (F : _ →+ R).comp (of A i) = (G : _ →+ R).comp (of A i)) : F = G :=
ring_hom.coe_add_monoid_hom_injective $ direct_sum.add_hom_ext' h

/-- Two `ring_hom`s out of a direct sum are equal if they agree on the generators. -/
lemma ring_hom_ext ⦃f g : (⨁ i, A i) →+* R⦄ (h : ∀ i x, f (of A i x) = g (of A i x)) :
f = g :=
ring_hom_ext' $ λ i, add_monoid_hom.ext $ h i

/-- A family of `add_monoid_hom`s preserving `direct_sum.ghas_one.one` and `direct_sum.ghas_mul.mul`
describes a `ring_hom`s on `⨁ i, A i`. This is a stronger version of `direct_sum.to_monoid`.
Expand Down Expand Up @@ -475,15 +480,6 @@ 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

0 comments on commit e5ba338

Please sign in to comment.