Skip to content

Commit

Permalink
chore(Data/Finsupp): split off extensionality from Defs.lean
Browse files Browse the repository at this point in the history
These results depend on `Submonoid`, while nothing else in `Finsupp` does. So let's move them to their own file.
  • Loading branch information
Vierkantor committed Nov 25, 2024
1 parent c4f3076 commit 890c963
Show file tree
Hide file tree
Showing 4 changed files with 74 additions and 45 deletions.
1 change: 1 addition & 0 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2444,6 +2444,7 @@ import Mathlib.Data.Finsupp.Basic
import Mathlib.Data.Finsupp.BigOperators
import Mathlib.Data.Finsupp.Defs
import Mathlib.Data.Finsupp.Encodable
import Mathlib.Data.Finsupp.Ext
import Mathlib.Data.Finsupp.Fin
import Mathlib.Data.Finsupp.Fintype
import Mathlib.Data.Finsupp.Indicator
Expand Down
1 change: 1 addition & 0 deletions Mathlib/Algebra/BigOperators/Finsupp.lean
Original file line number Diff line number Diff line change
Expand Up @@ -7,6 +7,7 @@ import Mathlib.Algebra.BigOperators.Pi
import Mathlib.Algebra.BigOperators.Ring
import Mathlib.Algebra.BigOperators.Fin
import Mathlib.Algebra.Group.Submonoid.Membership
import Mathlib.Data.Finsupp.Ext
import Mathlib.Data.Finsupp.Fin
import Mathlib.Data.Finsupp.Indicator

Expand Down
46 changes: 1 addition & 45 deletions Mathlib/Data/Finsupp/Defs.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johannes Hölzl, Kim Morrison
-/
import Mathlib.Algebra.Group.Indicator
import Mathlib.Algebra.Group.Submonoid.Basic
import Mathlib.Data.Finset.Max
import Mathlib.Data.Set.Finite.Basic
import Mathlib.Algebra.Group.TypeTags.Hom

Check warning on line 9 in Mathlib/Data/Finsupp/Defs.lean

View workflow job for this annotation

GitHub Actions / Build

unused import (use `lake exe shake --fix` to fix this, or `lake exe shake --update` to ignore)
Expand Down Expand Up @@ -81,6 +80,7 @@ This file is a `noncomputable theory` and uses classical logic throughout.
-/

assert_not_exists Submonoid

noncomputable section

Expand Down Expand Up @@ -1130,50 +1130,6 @@ theorem induction_on_min₂ (f : α →₀ M) (h0 : p 0)

end LinearOrder

@[simp]
theorem add_closure_setOf_eq_single :
AddSubmonoid.closure { f : α →₀ M | ∃ a b, f = single a b } = ⊤ :=
top_unique fun x _hx =>
Finsupp.induction x (AddSubmonoid.zero_mem _) fun a b _f _ha _hb hf =>
AddSubmonoid.add_mem _ (AddSubmonoid.subset_closure <| ⟨a, b, rfl⟩) hf

/-- If two additive homomorphisms from `α →₀ M` are equal on each `single a b`,
then they are equal. -/
theorem addHom_ext [AddZeroClass N] ⦃f g : (α →₀ M) →+ N⦄
(H : ∀ x y, f (single x y) = g (single x y)) : f = g := by
refine AddMonoidHom.eq_of_eqOn_denseM add_closure_setOf_eq_single ?_
rintro _ ⟨x, y, rfl⟩
apply H

/-- If two additive homomorphisms from `α →₀ M` are equal on each `single a b`,
then they are equal.
We formulate this using equality of `AddMonoidHom`s so that `ext` tactic can apply a type-specific
extensionality lemma after this one. E.g., if the fiber `M` is `ℕ` or `ℤ`, then it suffices to
verify `f (single a 1) = g (single a 1)`. -/
@[ext high]
theorem addHom_ext' [AddZeroClass N] ⦃f g : (α →₀ M) →+ N⦄
(H : ∀ x, f.comp (singleAddHom x) = g.comp (singleAddHom x)) : f = g :=
addHom_ext fun x => DFunLike.congr_fun (H x)

theorem mulHom_ext [MulOneClass N] ⦃f g : Multiplicative (α →₀ M) →* N⦄
(H : ∀ x y, f (Multiplicative.ofAdd <| single x y) = g (Multiplicative.ofAdd <| single x y)) :
f = g :=
MonoidHom.ext <|
DFunLike.congr_fun <| by
have := @addHom_ext α M (Additive N) _ _
(MonoidHom.toAdditive'' f) (MonoidHom.toAdditive'' g) H
ext
rw [DFunLike.ext_iff] at this
apply this

@[ext]
theorem mulHom_ext' [MulOneClass N] {f g : Multiplicative (α →₀ M) →* N}
(H : ∀ x, f.comp (AddMonoidHom.toMultiplicative (singleAddHom x)) =
g.comp (AddMonoidHom.toMultiplicative (singleAddHom x))) :
f = g :=
mulHom_ext fun x => DFunLike.congr_fun (H x)

theorem mapRange_add [AddZeroClass N] {f : M → N} {hf : f 0 = 0}
(hf' : ∀ x y, f (x + y) = f x + f y) (v₁ v₂ : α →₀ M) :
mapRange f hf (v₁ + v₂) = mapRange f hf v₁ + mapRange f hf v₂ :=
Expand Down
71 changes: 71 additions & 0 deletions Mathlib/Data/Finsupp/Ext.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,71 @@
/-
Copyright (c) 2017 Johannes Hölzl. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johannes Hölzl, Kim Morrison
-/
import Mathlib.Algebra.Group.Submonoid.Basic
import Mathlib.Data.Finsupp.Defs

/-!
# Extensionality for maps on `Finsupp`
This file contains some extensionality principles for maps on `Finsupp`.
These have been moved to their own file to avoid depending on submonoids when defining `Finsupp`.
## Main results
* `Finsupp.add_closure_setOf_eq_single`: `Finsupp` is generated by all the `single`s
* `Finsupp.addHom_ext`: additive homomorphisms that are equal on each `single` are equal everywhere
-/

variable {α M N : Type*}

namespace Finsupp

variable [AddZeroClass M]

@[simp]
theorem add_closure_setOf_eq_single :
AddSubmonoid.closure { f : α →₀ M | ∃ a b, f = single a b } = ⊤ :=
top_unique fun x _hx =>
Finsupp.induction x (AddSubmonoid.zero_mem _) fun a b _f _ha _hb hf =>
AddSubmonoid.add_mem _ (AddSubmonoid.subset_closure <| ⟨a, b, rfl⟩) hf

/-- If two additive homomorphisms from `α →₀ M` are equal on each `single a b`,
then they are equal. -/
theorem addHom_ext [AddZeroClass N] ⦃f g : (α →₀ M) →+ N⦄
(H : ∀ x y, f (single x y) = g (single x y)) : f = g := by
refine AddMonoidHom.eq_of_eqOn_denseM add_closure_setOf_eq_single ?_
rintro _ ⟨x, y, rfl⟩
apply H

/-- If two additive homomorphisms from `α →₀ M` are equal on each `single a b`,
then they are equal.
We formulate this using equality of `AddMonoidHom`s so that `ext` tactic can apply a type-specific
extensionality lemma after this one. E.g., if the fiber `M` is `ℕ` or `ℤ`, then it suffices to
verify `f (single a 1) = g (single a 1)`. -/
@[ext high]
theorem addHom_ext' [AddZeroClass N] ⦃f g : (α →₀ M) →+ N⦄
(H : ∀ x, f.comp (singleAddHom x) = g.comp (singleAddHom x)) : f = g :=
addHom_ext fun x => DFunLike.congr_fun (H x)

theorem mulHom_ext [MulOneClass N] ⦃f g : Multiplicative (α →₀ M) →* N⦄
(H : ∀ x y, f (Multiplicative.ofAdd <| single x y) = g (Multiplicative.ofAdd <| single x y)) :
f = g :=
MonoidHom.ext <|
DFunLike.congr_fun <| by
have := @addHom_ext α M (Additive N) _ _
(MonoidHom.toAdditive'' f) (MonoidHom.toAdditive'' g) H
ext
rw [DFunLike.ext_iff] at this
apply this

@[ext]
theorem mulHom_ext' [MulOneClass N] {f g : Multiplicative (α →₀ M) →* N}
(H : ∀ x, f.comp (AddMonoidHom.toMultiplicative (singleAddHom x)) =
g.comp (AddMonoidHom.toMultiplicative (singleAddHom x))) :
f = g :=
mulHom_ext fun x => DFunLike.congr_fun (H x)

end Finsupp

0 comments on commit 890c963

Please sign in to comment.