mathlib documentation

data.​finsupp.​basic

data.​finsupp.​basic

Type of functions with finite support

For any type α and a type β with zero, we define the type finsupp α β of finitely supported functions from α to β, i.e. the functions which are zero everywhere on α except on a finite set. We write this in infix notation as α →₀ β.

Functions with finite support provide the basis for the following concrete instances:

Most of the theory assumes that the range is a commutative monoid. This gives us the big sum operator as a powerful way to construct finsupp elements.

A general piece of advice is to not use α →₀ β directly, as the type class setup might not be a good fit. Defining a copy and selecting the instances that are best suited for the application works better.

Implementation notes

This file is a noncomputable theory and uses classical logic throughout.

Notation

This file defines α →₀ β as notation for finsupp α β.

structure finsupp (α : Type u_10) (β : Type u_11) [has_zero β] :
Type (max u_10 u_11)

finsupp α β, denoted α →₀ β, is the type of functions f : α → β such that f x = 0 for all but finitely many x.

Basic declarations about finsupp

@[instance]
def finsupp.​has_coe_to_fun {α : Type u_1} {β : Type u_2} [has_zero β] :

Equations
@[instance]
def finsupp.​has_zero {α : Type u_1} {β : Type u_2} [has_zero β] :

Equations
@[simp]
theorem finsupp.​zero_apply {α : Type u_1} {β : Type u_2} [has_zero β] {a : α} :
0 a = 0

@[simp]
theorem finsupp.​support_zero {α : Type u_1} {β : Type u_2} [has_zero β] :

@[instance]
def finsupp.​inhabited {α : Type u_1} {β : Type u_2} [has_zero β] :

Equations
@[simp]
theorem finsupp.​mem_support_iff {α : Type u_1} {β : Type u_2} [has_zero β] {f : α →₀ β} {a : α} :
a f.support f a 0

theorem finsupp.​not_mem_support_iff {α : Type u_1} {β : Type u_2} [has_zero β] {f : α →₀ β} {a : α} :
a f.support f a = 0

@[ext]
theorem finsupp.​ext {α : Type u_1} {β : Type u_2} [has_zero β] {f g : α →₀ β} :
(∀ (a : α), f a = g a)f = g

theorem finsupp.​ext_iff {α : Type u_1} {β : Type u_2} [has_zero β] {f g : α →₀ β} :
f = g ∀ (a : α), f a = g a

@[simp]
theorem finsupp.​support_eq_empty {α : Type u_1} {β : Type u_2} [has_zero β] {f : α →₀ β} :
f.support = f = 0

@[instance]
def finsupp.​finsupp.​decidable_eq {α : Type u_1} {β : Type u_2} [has_zero β] [decidable_eq α] [decidable_eq β] :

Equations
theorem finsupp.​finite_supp {α : Type u_1} {β : Type u_2} [has_zero β] (f : α →₀ β) :
{a : α | f a 0}.finite

theorem finsupp.​support_subset_iff {α : Type u_1} {β : Type u_2} [has_zero β] {s : set α} {f : α →₀ β} :
(f.support) s ∀ (a : α), a sf a = 0

def finsupp.​equiv_fun_on_fintype {α : Type u_1} {β : Type u_2} [has_zero β] [fintype α] :
→₀ β) (α → β)

Given fintype α, equiv_fun_on_fintype is the equiv between α →₀ β and α → β. (All functions on a finite type are finitely supported.)

Equations

Declarations about single

def finsupp.​single {α : Type u_1} {β : Type u_2} [has_zero β] :
α → β → α →₀ β

single a b is the finitely supported function which has value b at a and zero otherwise.

Equations
theorem finsupp.​single_apply {α : Type u_1} {β : Type u_2} [has_zero β] {a a' : α} {b : β} :
(finsupp.single a b) a' = ite (a = a') b 0

@[simp]
theorem finsupp.​single_eq_same {α : Type u_1} {β : Type u_2} [has_zero β] {a : α} {b : β} :

@[simp]
theorem finsupp.​single_eq_of_ne {α : Type u_1} {β : Type u_2} [has_zero β] {a a' : α} {b : β} :
a a'(finsupp.single a b) a' = 0

@[simp]
theorem finsupp.​single_zero {α : Type u_1} {β : Type u_2} [has_zero β] {a : α} :

theorem finsupp.​support_single_ne_zero {α : Type u_1} {β : Type u_2} [has_zero β] {a : α} {b : β} :
b 0(finsupp.single a b).support = {a}

theorem finsupp.​support_single_subset {α : Type u_1} {β : Type u_2} [has_zero β] {a : α} {b : β} :

theorem finsupp.​single_injective {α : Type u_1} {β : Type u_2} [has_zero β] (a : α) :

theorem finsupp.​single_eq_single_iff {α : Type u_1} {β : Type u_2} [has_zero β] (a₁ a₂ : α) (b₁ b₂ : β) :
finsupp.single a₁ b₁ = finsupp.single a₂ b₂ a₁ = a₂ b₁ = b₂ b₁ = 0 b₂ = 0

theorem finsupp.​single_left_inj {α : Type u_1} {β : Type u_2} [has_zero β] {a a' : α} {b : β} :
b 0(finsupp.single a b = finsupp.single a' b a = a')

theorem finsupp.​single_eq_zero {α : Type u_1} {β : Type u_2} [has_zero β] {a : α} {b : β} :
finsupp.single a b = 0 b = 0

theorem finsupp.​single_swap {α : Type u_1} {β : Type u_2} [has_zero β] (a₁ a₂ : α) (b : β) :
(finsupp.single a₁ b) a₂ = (finsupp.single a₂ b) a₁

theorem finsupp.​unique_single {α : Type u_1} {β : Type u_2} [has_zero β] [unique α] (x : α →₀ β) :

@[simp]
theorem finsupp.​unique_single_eq_iff {α : Type u_1} {β : Type u_2} [has_zero β] {a a' : α} {b : β} [unique α] {b' : β} :

Declarations about on_finset

def finsupp.​on_finset {α : Type u_1} {β : Type u_2} [has_zero β] (s : finset α) (f : α → β) :
(∀ (a : α), f a 0a s)α →₀ β

on_finset s f hf is the finsupp function representing f restricted to the finset s. The function needs to be 0 outside of s. Use this when the set needs to be filtered anyways, otherwise a better set representation is often available.

Equations
@[simp]
theorem finsupp.​on_finset_apply {α : Type u_1} {β : Type u_2} [has_zero β] {s : finset α} {f : α → β} {hf : ∀ (a : α), f a 0a s} {a : α} :
(finsupp.on_finset s f hf) a = f a

@[simp]
theorem finsupp.​support_on_finset_subset {α : Type u_1} {β : Type u_2} [has_zero β] {s : finset α} {f : α → β} {hf : ∀ (a : α), f a 0a s} :

Declarations about map_range

def finsupp.​map_range {α : Type u_1} {β₁ : Type u_8} {β₂ : Type u_9} [has_zero β₁] [has_zero β₂] (f : β₁ → β₂) :
f 0 = 0→₀ β₁)α →₀ β₂

The composition of f : β₁ → β₂ and g : α →₀ β₁ is map_range f hf g : α →₀ β₂, well-defined when f 0 = 0.

Equations
@[simp]
theorem finsupp.​map_range_apply {α : Type u_1} {β₁ : Type u_8} {β₂ : Type u_9} [has_zero β₁] [has_zero β₂] {f : β₁ → β₂} {hf : f 0 = 0} {g : α →₀ β₁} {a : α} :
(finsupp.map_range f hf g) a = f (g a)

@[simp]
theorem finsupp.​map_range_zero {α : Type u_1} {β₁ : Type u_8} {β₂ : Type u_9} [has_zero β₁] [has_zero β₂] {f : β₁ → β₂} {hf : f 0 = 0} :

theorem finsupp.​support_map_range {α : Type u_1} {β₁ : Type u_8} {β₂ : Type u_9} [has_zero β₁] [has_zero β₂] {f : β₁ → β₂} {hf : f 0 = 0} {g : α →₀ β₁} :

@[simp]
theorem finsupp.​map_range_single {α : Type u_1} {β₁ : Type u_8} {β₂ : Type u_9} [has_zero β₁] [has_zero β₂] {f : β₁ → β₂} {hf : f 0 = 0} {a : α} {b : β₁} :

Declarations about emb_domain

def finsupp.​emb_domain {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [has_zero β] :
(α₁ α₂)(α₁ →₀ β)α₂ →₀ β

Given f : α₁ ↪ α₂ and v : α₁ →₀ β, emb_domain f v : α₂ →₀ β is the finitely supported function whose value at f a : α₂ is v a. For a b : α₂ outside the range of f, it is zero.

Equations
theorem finsupp.​support_emb_domain {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [has_zero β] (f : α₁ α₂) (v : α₁ →₀ β) :

theorem finsupp.​emb_domain_zero {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [has_zero β] (f : α₁ α₂) :

theorem finsupp.​emb_domain_apply {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [has_zero β] (f : α₁ α₂) (v : α₁ →₀ β) (a : α₁) :

theorem finsupp.​emb_domain_notin_range {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [has_zero β] (f : α₁ α₂) (v : α₁ →₀ β) (a : α₂) :

theorem finsupp.​emb_domain_inj {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [has_zero β] {f : α₁ α₂} {l₁ l₂ : α₁ →₀ β} :
finsupp.emb_domain f l₁ = finsupp.emb_domain f l₂ l₁ = l₂

theorem finsupp.​emb_domain_map_range {α₁ : Type u_6} {α₂ : Type u_7} {β₁ : Type u_1} {β₂ : Type u_2} [has_zero β₁] [has_zero β₂] (f : α₁ α₂) (g : β₁ → β₂) (p : α₁ →₀ β₁) (hg : g 0 = 0) :

theorem finsupp.​single_of_emb_domain_single {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [has_zero β] (l : α₁ →₀ β) (f : α₁ α₂) (a : α₂) (b : β) :
b 0finsupp.emb_domain f l = finsupp.single a b(∃ (x : α₁), l = finsupp.single x b f x = a)

Declarations about zip_with

def finsupp.​zip_with {α : Type u_1} {β : Type u_2} {β₁ : Type u_8} {β₂ : Type u_9} [has_zero β] [has_zero β₁] [has_zero β₂] (f : β₁ → β₂ → β) :
f 0 0 = 0→₀ β₁)→₀ β₂)α →₀ β

zip_with f hf g₁ g₂ is the finitely supported function satisfying zip_with f hf g₁ g₂ a = f (g₁ a) (g₂ a), and it is well-defined when f 0 0 = 0.

Equations
@[simp]
theorem finsupp.​zip_with_apply {α : Type u_1} {β : Type u_2} {β₁ : Type u_8} {β₂ : Type u_9} [has_zero β] [has_zero β₁] [has_zero β₂] {f : β₁ → β₂ → β} {hf : f 0 0 = 0} {g₁ : α →₀ β₁} {g₂ : α →₀ β₂} {a : α} :
(finsupp.zip_with f hf g₁ g₂) a = f (g₁ a) (g₂ a)

theorem finsupp.​support_zip_with {α : Type u_1} {β : Type u_2} {β₁ : Type u_8} {β₂ : Type u_9} [has_zero β] [has_zero β₁] [has_zero β₂] {f : β₁ → β₂ → β} {hf : f 0 0 = 0} {g₁ : α →₀ β₁} {g₂ : α →₀ β₂} :
(finsupp.zip_with f hf g₁ g₂).support g₁.support g₂.support

Declarations about erase

def finsupp.​erase {α : Type u_1} {β : Type u_2} [has_zero β] :
α → →₀ β)α →₀ β

erase a f is the finitely supported function equal to f except at a where it is equal to 0.

Equations
@[simp]
theorem finsupp.​support_erase {α : Type u_1} {β : Type u_2} [has_zero β] {a : α} {f : α →₀ β} :

@[simp]
theorem finsupp.​erase_same {α : Type u_1} {β : Type u_2} [has_zero β] {a : α} {f : α →₀ β} :

@[simp]
theorem finsupp.​erase_ne {α : Type u_1} {β : Type u_2} [has_zero β] {a a' : α} {f : α →₀ β} :
a' a(finsupp.erase a f) a' = f a'

@[simp]
theorem finsupp.​erase_single {α : Type u_1} {β : Type u_2} [has_zero β] {a : α} {b : β} :

theorem finsupp.​erase_single_ne {α : Type u_1} {β : Type u_2} [has_zero β] {a a' : α} {b : β} :

Declarations about sum and prod

In most of this section, the domain β is assumed to be an add_monoid.

def finsupp.​sum {α : Type u_1} {β : Type u_2} {γ : Type u_3} [has_zero β] [add_comm_monoid γ] :
→₀ β)(α → β → γ) → γ

sum f g is the sum of g a (f a) over the support of f.

Equations
def finsupp.​prod {α : Type u_1} {β : Type u_2} {γ : Type u_3} [has_zero β] [comm_monoid γ] :
→₀ β)(α → β → γ) → γ

prod f g is the product of g a (f a) over the support of f.

Equations
theorem finsupp.​prod_fintype {α : Type u_1} {β : Type u_2} {γ : Type u_3} [fintype α] [has_zero β] [comm_monoid γ] (f : α →₀ β) (g : α → β → γ) :
(∀ (i : α), g i 0 = 1)f.prod g = finset.univ.prod (λ (i : α), g i (f i))

theorem finsupp.​sum_fintype {α : Type u_1} {β : Type u_2} {γ : Type u_3} [fintype α] [has_zero β] [add_comm_monoid γ] (f : α →₀ β) (g : α → β → γ) :
(∀ (i : α), g i 0 = 0)f.sum g = finset.univ.sum (λ (i : α), g i (f i))

theorem finsupp.​prod_map_range_index {α : Type u_1} {γ : Type u_3} {β₁ : Type u_8} {β₂ : Type u_9} [has_zero β₁] [has_zero β₂] [comm_monoid γ] {f : β₁ → β₂} {hf : f 0 = 0} {g : α →₀ β₁} {h : α → β₂ → γ} :
(∀ (a : α), h a 0 = 1)(finsupp.map_range f hf g).prod h = g.prod (λ (a : α) (b : β₁), h a (f b))

theorem finsupp.​sum_map_range_index {α : Type u_1} {γ : Type u_3} {β₁ : Type u_8} {β₂ : Type u_9} [has_zero β₁] [has_zero β₂] [add_comm_monoid γ] {f : β₁ → β₂} {hf : f 0 = 0} {g : α →₀ β₁} {h : α → β₂ → γ} :
(∀ (a : α), h a 0 = 0)(finsupp.map_range f hf g).sum h = g.sum (λ (a : α) (b : β₁), h a (f b))

theorem finsupp.​sum_zero_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid β] [add_comm_monoid γ] {h : α → β → γ} :
0.sum h = 0

theorem finsupp.​prod_zero_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid β] [comm_monoid γ] {h : α → β → γ} :
0.prod h = 1

theorem finsupp.​prod_comm {α : Type u_1} {β : Type u_2} {γ : Type u_3} {α' : Type u_4} [has_zero β] {β' : Type u_5} [has_zero β'] (f : α →₀ β) (g : α' →₀ β') [comm_monoid γ] (h : α → β → α' → β' → γ) :
f.prod (λ (x : α) (v : β), g.prod (λ (x' : α') (v' : β'), h x v x' v')) = g.prod (λ (x' : α') (v' : β'), f.prod (λ (x : α) (v : β), h x v x' v'))

theorem finsupp.​sum_comm {α : Type u_1} {β : Type u_2} {γ : Type u_3} {α' : Type u_4} [has_zero β] {β' : Type u_5} [has_zero β'] (f : α →₀ β) (g : α' →₀ β') [add_comm_monoid γ] (h : α → β → α' → β' → γ) :
f.sum (λ (x : α) (v : β), g.sum (λ (x' : α') (v' : β'), h x v x' v')) = g.sum (λ (x' : α') (v' : β'), f.sum (λ (x : α) (v : β), h x v x' v'))

@[simp]
theorem finsupp.​prod_ite_eq {α : Type u_1} {β : Type u_2} {γ : Type u_3} [has_zero β] [comm_monoid γ] (f : α →₀ β) (a : α) (b : α → β → γ) :
f.prod (λ (x : α) (v : β), ite (a = x) (b x v) 1) = ite (a f.support) (b a (f a)) 1

@[simp]
theorem finsupp.​sum_ite_eq {α : Type u_1} {β : Type u_2} {γ : Type u_3} [has_zero β] [add_comm_monoid γ] (f : α →₀ β) (a : α) (b : α → β → γ) :
f.sum (λ (x : α) (v : β), ite (a = x) (b x v) 0) = ite (a f.support) (b a (f a)) 0

@[simp]
theorem finsupp.​prod_ite_eq' {α : Type u_1} {β : Type u_2} {γ : Type u_3} [has_zero β] [comm_monoid γ] (f : α →₀ β) (a : α) (b : α → β → γ) :
f.prod (λ (x : α) (v : β), ite (x = a) (b x v) 1) = ite (a f.support) (b a (f a)) 1

A restatement of prod_ite_eq with the equality test reversed.

@[simp]
theorem finsupp.​sum_ite_eq' {α : Type u_1} {β : Type u_2} {γ : Type u_3} [has_zero β] [add_comm_monoid γ] (f : α →₀ β) (a : α) (b : α → β → γ) :
f.sum (λ (x : α) (v : β), ite (x = a) (b x v) 0) = ite (a f.support) (b a (f a)) 0

A restatement of sum_ite_eq with the equality test reversed.

@[simp]
theorem finsupp.​prod_pow {α : Type u_1} {γ : Type u_3} [fintype α] [comm_monoid γ] (f : α →₀ ) (g : α → γ) :
f.prod (λ (a : α) (b : ), g a ^ b) = finset.univ.prod (λ (a : α), g a ^ f a)

theorem finsupp.​on_finset_sum {α : Type u_1} {β : Type u_2} {γ : Type u_3} [has_zero β] [add_comm_monoid γ] {s : finset α} {f : α → β} {g : α → β → γ} (hf : ∀ (a : α), f a 0a s) :
(∀ (a : α), g a 0 = 0)(finsupp.on_finset s f hf).sum g = s.sum (λ (a : α), g a (f a))

If g maps a second argument of 0 to 0, summing it over the result of on_finset is the same as summing it over the original finset.

theorem finsupp.​sum_single_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_monoid β] [add_comm_monoid γ] {a : α} {b : β} {h : α → β → γ} :
h a 0 = 0(finsupp.single a b).sum h = h a b

theorem finsupp.​prod_single_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_monoid β] [comm_monoid γ] {a : α} {b : β} {h : α → β → γ} :
h a 0 = 1(finsupp.single a b).prod h = h a b

@[instance]
def finsupp.​has_add {α : Type u_1} {β : Type u_2} [add_monoid β] :
has_add →₀ β)

Equations
@[simp]
theorem finsupp.​add_apply {α : Type u_1} {β : Type u_2} [add_monoid β] {g₁ g₂ : α →₀ β} {a : α} :
(g₁ + g₂) a = g₁ a + g₂ a

theorem finsupp.​support_add {α : Type u_1} {β : Type u_2} [add_monoid β] {g₁ g₂ : α →₀ β} :
(g₁ + g₂).support g₁.support g₂.support

theorem finsupp.​support_add_eq {α : Type u_1} {β : Type u_2} [add_monoid β] {g₁ g₂ : α →₀ β} :
disjoint g₁.support g₂.support(g₁ + g₂).support = g₁.support g₂.support

@[simp]
theorem finsupp.​single_add {α : Type u_1} {β : Type u_2} [add_monoid β] {a : α} {b₁ b₂ : β} :
finsupp.single a (b₁ + b₂) = finsupp.single a b₁ + finsupp.single a b₂

@[instance]
def finsupp.​add_monoid {α : Type u_1} {β : Type u_2} [add_monoid β] :

Equations
def finsupp.​eval_add_hom {α : Type u_1} {β : Type u_2} [add_monoid β] :
α → →₀ β) →+ β

Evaluation of a function f : α →₀ β at a point as an additive monoid homomorphism.

Equations
@[simp]
theorem finsupp.​eval_add_hom_apply {α : Type u_1} {β : Type u_2} [add_monoid β] (a : α) (g : α →₀ β) :

theorem finsupp.​single_add_erase {α : Type u_1} {β : Type u_2} [add_monoid β] {a : α} {f : α →₀ β} :

theorem finsupp.​erase_add_single {α : Type u_1} {β : Type u_2} [add_monoid β] {a : α} {f : α →₀ β} :

@[simp]
theorem finsupp.​erase_add {α : Type u_1} {β : Type u_2} [add_monoid β] (a : α) (f f' : α →₀ β) :

theorem finsupp.​induction {α : Type u_1} {β : Type u_2} [add_monoid β] {p : →₀ β) → Prop} (f : α →₀ β) :
p 0(∀ (a : α) (b : β) (f : α →₀ β), a f.supportb 0p fp (finsupp.single a b + f))p f

theorem finsupp.​induction₂ {α : Type u_1} {β : Type u_2} [add_monoid β] {p : →₀ β) → Prop} (f : α →₀ β) :
p 0(∀ (a : α) (b : β) (f : α →₀ β), a f.supportb 0p fp (f + finsupp.single a b))p f

theorem finsupp.​map_range_add {α : Type u_1} {β₁ : Type u_8} {β₂ : Type u_9} [add_monoid β₁] [add_monoid β₂] {f : β₁ → β₂} {hf : f 0 = 0} (hf' : ∀ (x y : β₁), f (x + y) = f x + f y) (v₁ v₂ : α →₀ β₁) :
finsupp.map_range f hf (v₁ + v₂) = finsupp.map_range f hf v₁ + finsupp.map_range f hf v₂

@[instance]
def finsupp.​nat_sub {α : Type u_1} :

Equations
@[simp]
theorem finsupp.​nat_sub_apply {α : Type u_1} {g₁ g₂ : α →₀ } {a : α} :
(g₁ - g₂) a = g₁ a - g₂ a

@[simp]
theorem finsupp.​single_sub {α : Type u_1} {a : α} {n₁ n₂ : } :
finsupp.single a (n₁ - n₂) = finsupp.single a n₁ - finsupp.single a n₂

theorem finsupp.​sub_single_one_add {α : Type u_1} {a : α} {u u' : α →₀ } :
u a 0u - finsupp.single a 1 + u' = u + u' - finsupp.single a 1

theorem finsupp.​add_sub_single_one {α : Type u_1} {a : α} {u u' : α →₀ } :
u' a 0u + (u' - finsupp.single a 1) = u + u' - finsupp.single a 1

theorem finsupp.​single_multiset_sum {α : Type u_1} {β : Type u_2} [add_comm_monoid β] (s : multiset β) (a : α) :

theorem finsupp.​single_finset_sum {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid β] (s : finset γ) (f : γ → β) (a : α) :
finsupp.single a (s.sum (λ (b : γ), f b)) = s.sum (λ (b : γ), finsupp.single a (f b))

theorem finsupp.​single_sum {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [has_zero γ] [add_comm_monoid β] (s : δ →₀ γ) (f : δ → γ → β) (a : α) :
finsupp.single a (s.sum f) = s.sum (λ (d : δ) (c : γ), finsupp.single a (f d c))

theorem finsupp.​sum_neg_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_group β] [add_comm_monoid γ] {g : α →₀ β} {h : α → β → γ} :
(∀ (a : α), h a 0 = 0)(-g).sum h = g.sum (λ (a : α) (b : β), h a (-b))

theorem finsupp.​prod_neg_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_group β] [comm_monoid γ] {g : α →₀ β} {h : α → β → γ} :
(∀ (a : α), h a 0 = 1)(-g).prod h = g.prod (λ (a : α) (b : β), h a (-b))

@[simp]
theorem finsupp.​neg_apply {α : Type u_1} {β : Type u_2} [add_group β] {g : α →₀ β} {a : α} :
(-g) a = -g a

@[simp]
theorem finsupp.​sub_apply {α : Type u_1} {β : Type u_2} [add_group β] {g₁ g₂ : α →₀ β} {a : α} :
(g₁ - g₂) a = g₁ a - g₂ a

@[simp]
theorem finsupp.​support_neg {α : Type u_1} {β : Type u_2} [add_group β] {f : α →₀ β} :

@[simp]
theorem finsupp.​sum_apply {α : Type u_1} {β : Type u_2} {α₁ : Type u_6} {β₁ : Type u_8} [has_zero β₁] [add_comm_monoid β] {f : α₁ →₀ β₁} {g : α₁ → β₁ → α →₀ β} {a₂ : α} :
(f.sum g) a₂ = f.sum (λ (a₁ : α₁) (b : β₁), (g a₁ b) a₂)

theorem finsupp.​support_sum {α : Type u_1} {β : Type u_2} {α₁ : Type u_6} {β₁ : Type u_8} [has_zero β₁] [add_comm_monoid β] {f : α₁ →₀ β₁} {g : α₁ → β₁ → α →₀ β} :
(f.sum g).support f.support.bind (λ (a : α₁), (g a (f a)).support)

@[simp]
theorem finsupp.​sum_zero {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid β] [add_comm_monoid γ] {f : α →₀ β} :
f.sum (λ (a : α) (b : β), 0) = 0

@[simp]
theorem finsupp.​sum_add {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid β] [add_comm_monoid γ] {f : α →₀ β} {h₁ h₂ : α → β → γ} :
f.sum (λ (a : α) (b : β), h₁ a b + h₂ a b) = f.sum h₁ + f.sum h₂

@[simp]
theorem finsupp.​sum_neg {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid β] [add_comm_group γ] {f : α →₀ β} {h : α → β → γ} :
f.sum (λ (a : α) (b : β), -h a b) = -f.sum h

@[simp]
theorem finsupp.​sum_sub {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid β] [add_comm_group γ] {f : α →₀ β} {h₁ h₂ : α → β → γ} :
f.sum (λ (a : α) (b : β), h₁ a b - h₂ a b) = f.sum h₁ - f.sum h₂

@[simp]
theorem finsupp.​sum_single {α : Type u_1} {β : Type u_2} [add_comm_monoid β] (f : α →₀ β) :

theorem finsupp.​prod_add_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid β] [comm_monoid γ] {f g : α →₀ β} {h : α → β → γ} :
(∀ (a : α), h a 0 = 1)(∀ (a : α) (b₁ b₂ : β), h a (b₁ + b₂) = h a b₁ * h a b₂)(f + g).prod h = f.prod h * g.prod h

theorem finsupp.​sum_add_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid β] [add_comm_monoid γ] {f g : α →₀ β} {h : α → β → γ} :
(∀ (a : α), h a 0 = 0)(∀ (a : α) (b₁ b₂ : β), h a (b₁ + b₂) = h a b₁ + h a b₂)(f + g).sum h = f.sum h + g.sum h

theorem finsupp.​sum_sub_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_group β] [add_comm_group γ] {f g : α →₀ β} {h : α → β → γ} :
(∀ (a : α) (b₁ b₂ : β), h a (b₁ - b₂) = h a b₁ - h a b₂)(f - g).sum h = f.sum h - g.sum h

theorem finsupp.​prod_finset_sum_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} {ι : Type u_5} [add_comm_monoid β] [comm_monoid γ] {s : finset ι} {g : ι → α →₀ β} {h : α → β → γ} :
(∀ (a : α), h a 0 = 1)(∀ (a : α) (b₁ b₂ : β), h a (b₁ + b₂) = h a b₁ * h a b₂)s.prod (λ (i : ι), (g i).prod h) = (s.sum (λ (i : ι), g i)).prod h

theorem finsupp.​sum_finset_sum_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} {ι : Type u_5} [add_comm_monoid β] [add_comm_monoid γ] {s : finset ι} {g : ι → α →₀ β} {h : α → β → γ} :
(∀ (a : α), h a 0 = 0)(∀ (a : α) (b₁ b₂ : β), h a (b₁ + b₂) = h a b₁ + h a b₂)s.sum (λ (i : ι), (g i).sum h) = (s.sum (λ (i : ι), g i)).sum h

theorem finsupp.​sum_sum_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} {α₁ : Type u_6} {β₁ : Type u_8} [add_comm_monoid β₁] [add_comm_monoid β] [add_comm_monoid γ] {f : α₁ →₀ β₁} {g : α₁ → β₁ → α →₀ β} {h : α → β → γ} :
(∀ (a : α), h a 0 = 0)(∀ (a : α) (b₁ b₂ : β), h a (b₁ + b₂) = h a b₁ + h a b₂)(f.sum g).sum h = f.sum (λ (a : α₁) (b : β₁), (g a b).sum h)

theorem finsupp.​prod_sum_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} {α₁ : Type u_6} {β₁ : Type u_8} [add_comm_monoid β₁] [add_comm_monoid β] [comm_monoid γ] {f : α₁ →₀ β₁} {g : α₁ → β₁ → α →₀ β} {h : α → β → γ} :
(∀ (a : α), h a 0 = 1)(∀ (a : α) (b₁ b₂ : β), h a (b₁ + b₂) = h a b₁ * h a b₂)(f.sum g).prod h = f.prod (λ (a : α₁) (b : β₁), (g a b).prod h)

theorem finsupp.​multiset_sum_sum_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid β] [add_comm_monoid γ] (f : multiset →₀ β)) (h : α → β → γ) :
(∀ (a : α), h a 0 = 0)(∀ (a : α) (b₁ b₂ : β), h a (b₁ + b₂) = h a b₁ + h a b₂)f.sum.sum h = (multiset.map (λ (g : α →₀ β), g.sum h) f).sum

theorem finsupp.​multiset_map_sum {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [has_zero β] {f : α →₀ β} {m : γ → δ} {h : α → β → multiset γ} :
multiset.map m (f.sum h) = f.sum (λ (a : α) (b : β), multiset.map m (h a b))

theorem finsupp.​multiset_sum_sum {α : Type u_1} {β : Type u_2} {γ : Type u_3} [has_zero β] [add_comm_monoid γ] {f : α →₀ β} {h : α → β → multiset γ} :
(f.sum h).sum = f.sum (λ (a : α) (b : β), (h a b).sum)

def finsupp.​map_range.​add_monoid_hom {α : Type u_1} {β₁ : Type u_8} {β₂ : Type u_9} [add_comm_monoid β₁] [add_comm_monoid β₂] :
(β₁ →+ β₂)→₀ β₁) →+ α →₀ β₂

Composition with a fixed additive homomorphism is itself an additive homomorphism on functions.

Equations
theorem finsupp.​map_range_multiset_sum {α : Type u_1} {β₁ : Type u_8} {β₂ : Type u_9} [add_comm_monoid β₁] [add_comm_monoid β₂] (f : β₁ →+ β₂) (m : multiset →₀ β₁)) :

theorem finsupp.​map_range_finset_sum {α : Type u_1} {β₁ : Type u_8} {β₂ : Type u_9} [add_comm_monoid β₁] [add_comm_monoid β₂] (f : β₁ →+ β₂) {ι : Type u_2} (s : finset ι) (g : ι → α →₀ β₁) :
finsupp.map_range f _ (s.sum (λ (x : ι), g x)) = s.sum (λ (x : ι), finsupp.map_range f _ (g x))

Declarations about map_domain

def finsupp.​map_domain {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [add_comm_monoid β] :
(α₁ → α₂)(α₁ →₀ β)α₂ →₀ β

Given f : α₁ → α₂ and v : α₁ →₀ β, map_domain f v : α₂ →₀ β is the finitely supported function whose value at a : α₂ is the sum of v x over all x such that f x = a.

Equations
theorem finsupp.​map_domain_apply {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [add_comm_monoid β] {f : α₁ → α₂} (hf : function.injective f) (x : α₁ →₀ β) (a : α₁) :
(finsupp.map_domain f x) (f a) = x a

theorem finsupp.​map_domain_notin_range {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [add_comm_monoid β] {f : α₁ → α₂} (x : α₁ →₀ β) (a : α₂) :

theorem finsupp.​map_domain_id {α : Type u_1} {β : Type u_2} [add_comm_monoid β] {v : α →₀ β} :

theorem finsupp.​map_domain_comp {α : Type u_1} {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [add_comm_monoid β] {v : α →₀ β} {f : α → α₁} {g : α₁ → α₂} :

theorem finsupp.​map_domain_single {α : Type u_1} {β : Type u_2} {α₁ : Type u_6} [add_comm_monoid β] {f : α → α₁} {a : α} {b : β} :

@[simp]
theorem finsupp.​map_domain_zero {α : Type u_1} {β : Type u_2} {α₂ : Type u_7} [add_comm_monoid β] {f : α → α₂} :

theorem finsupp.​map_domain_congr {α : Type u_1} {β : Type u_2} {α₂ : Type u_7} [add_comm_monoid β] {v : α →₀ β} {f g : α → α₂} :
(∀ (x : α), x v.supportf x = g x)finsupp.map_domain f v = finsupp.map_domain g v

theorem finsupp.​map_domain_add {α : Type u_1} {β : Type u_2} {α₂ : Type u_7} [add_comm_monoid β] {v₁ v₂ : α →₀ β} {f : α → α₂} :

theorem finsupp.​map_domain_finset_sum {α : Type u_1} {β : Type u_2} {ι : Type u_5} {α₂ : Type u_7} [add_comm_monoid β] {f : α → α₂} {s : finset ι} {v : ι → α →₀ β} :
finsupp.map_domain f (s.sum (λ (i : ι), v i)) = s.sum (λ (i : ι), finsupp.map_domain f (v i))

theorem finsupp.​map_domain_sum {α : Type u_1} {β : Type u_2} {α₂ : Type u_7} {β₁ : Type u_8} [add_comm_monoid β] [has_zero β₁] {f : α → α₂} {s : α →₀ β₁} {v : α → β₁ → α →₀ β} :
finsupp.map_domain f (s.sum v) = s.sum (λ (a : α) (b : β₁), finsupp.map_domain f (v a b))

theorem finsupp.​map_domain_support {α : Type u_1} {β : Type u_2} {α₂ : Type u_7} [add_comm_monoid β] {f : α → α₂} {s : α →₀ β} :

theorem finsupp.​sum_map_domain_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} {α₂ : Type u_7} [add_comm_monoid β] [add_comm_monoid γ] {f : α → α₂} {s : α →₀ β} {h : α₂ → β → γ} :
(∀ (a : α₂), h a 0 = 0)(∀ (a : α₂) (b₁ b₂ : β), h a (b₁ + b₂) = h a b₁ + h a b₂)(finsupp.map_domain f s).sum h = s.sum (λ (a : α) (b : β), h (f a) b)

theorem finsupp.​prod_map_domain_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} {α₂ : Type u_7} [add_comm_monoid β] [comm_monoid γ] {f : α → α₂} {s : α →₀ β} {h : α₂ → β → γ} :
(∀ (a : α₂), h a 0 = 1)(∀ (a : α₂) (b₁ b₂ : β), h a (b₁ + b₂) = h a b₁ * h a b₂)(finsupp.map_domain f s).prod h = s.prod (λ (a : α) (b : β), h (f a) b)

theorem finsupp.​emb_domain_eq_map_domain {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [add_comm_monoid β] (f : α₁ α₂) (v : α₁ →₀ β) :

theorem finsupp.​map_domain_injective {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [add_comm_monoid β] {f : α₁ → α₂} :

Declarations about comap_domain

def finsupp.​comap_domain {α₁ : Type u_1} {α₂ : Type u_2} {γ : Type u_3} [has_zero γ] (f : α₁ → α₂) (l : α₂ →₀ γ) :
set.inj_on f (f ⁻¹' (l.support))α₁ →₀ γ

Given f : α₁ → α₂, l : α₂ →₀ γ and a proof hf that f is injective on the preimage of l.support, comap_domain f l hf is the finitely supported function from α₁ to γ given by composing l with f.

Equations
@[simp]
theorem finsupp.​comap_domain_apply {α₁ : Type u_1} {α₂ : Type u_2} {γ : Type u_3} [has_zero γ] (f : α₁ → α₂) (l : α₂ →₀ γ) (hf : set.inj_on f (f ⁻¹' (l.support))) (a : α₁) :
(finsupp.comap_domain f l hf) a = l (f a)

theorem finsupp.​sum_comap_domain {α₁ : Type u_1} {α₂ : Type u_2} {β : Type u_3} {γ : Type u_4} [has_zero β] [add_comm_monoid γ] (f : α₁ → α₂) (l : α₂ →₀ β) (g : α₂ → β → γ) (hf : set.bij_on f (f ⁻¹' (l.support)) (l.support)) :
(finsupp.comap_domain f l _).sum (g f) = l.sum g

theorem finsupp.​eq_zero_of_comap_domain_eq_zero {α₁ : Type u_1} {α₂ : Type u_2} {γ : Type u_3} [add_comm_monoid γ] (f : α₁ → α₂) (l : α₂ →₀ γ) (hf : set.bij_on f (f ⁻¹' (l.support)) (l.support)) :
finsupp.comap_domain f l _ = 0l = 0

theorem finsupp.​map_domain_comap_domain {α₁ : Type u_1} {α₂ : Type u_2} {γ : Type u_3} [add_comm_monoid γ] (f : α₁ → α₂) (l : α₂ →₀ γ) (hf : function.injective f) :

Declarations about filter

def finsupp.​filter {α : Type u_1} {β : Type u_2} [has_zero β] :
(α → Prop)→₀ β)α →₀ β

filter p f is the function which is f a if p a is true and 0 otherwise.

Equations
@[simp]
theorem finsupp.​filter_apply_pos {α : Type u_1} {β : Type u_2} [has_zero β] (p : α → Prop) (f : α →₀ β) {a : α} :
p a(finsupp.filter p f) a = f a

@[simp]
theorem finsupp.​filter_apply_neg {α : Type u_1} {β : Type u_2} [has_zero β] (p : α → Prop) (f : α →₀ β) {a : α} :
¬p a(finsupp.filter p f) a = 0

@[simp]
theorem finsupp.​support_filter {α : Type u_1} {β : Type u_2} [has_zero β] (p : α → Prop) (f : α →₀ β) :

theorem finsupp.​filter_zero {α : Type u_1} {β : Type u_2} [has_zero β] (p : α → Prop) :

@[simp]
theorem finsupp.​filter_single_of_pos {α : Type u_1} {β : Type u_2} [has_zero β] (p : α → Prop) {a : α} {b : β} :

@[simp]
theorem finsupp.​filter_single_of_neg {α : Type u_1} {β : Type u_2} [has_zero β] (p : α → Prop) {a : α} {b : β} :

theorem finsupp.​filter_pos_add_filter_neg {α : Type u_1} {β : Type u_2} [add_monoid β] (f : α →₀ β) (p : α → Prop) :
finsupp.filter p f + finsupp.filter (λ (a : α), ¬p a) f = f

Declarations about frange

def finsupp.​frange {α : Type u_1} {β : Type u_2} [has_zero β] :
→₀ β)finset β

frange f is the image of f on the support of f.

Equations
theorem finsupp.​mem_frange {α : Type u_1} {β : Type u_2} [has_zero β] {f : α →₀ β} {y : β} :
y f.frange y 0 ∃ (x : α), f x = y

theorem finsupp.​zero_not_mem_frange {α : Type u_1} {β : Type u_2} [has_zero β] {f : α →₀ β} :

theorem finsupp.​frange_single {α : Type u_1} {β : Type u_2} [has_zero β] {x : α} {y : β} :

Declarations about subtype_domain

def finsupp.​subtype_domain {α : Type u_1} {β : Type u_2} [has_zero β] (p : α → Prop) :
→₀ β)subtype p →₀ β

subtype_domain p f is the restriction of the finitely supported function f to the subtype p.

Equations
@[simp]
theorem finsupp.​support_subtype_domain {α : Type u_1} {β : Type u_2} {p : α → Prop} [has_zero β] {f : α →₀ β} :

@[simp]
theorem finsupp.​subtype_domain_apply {α : Type u_1} {β : Type u_2} {p : α → Prop} [has_zero β] {a : subtype p} {v : α →₀ β} :

@[simp]
theorem finsupp.​subtype_domain_zero {α : Type u_1} {β : Type u_2} {p : α → Prop} [has_zero β] :

theorem finsupp.​prod_subtype_domain_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} {p : α → Prop} [has_zero β] [comm_monoid γ] {v : α →₀ β} {h : α → β → γ} :
(∀ (x : α), x v.supportp x)(finsupp.subtype_domain p v).prod (λ (a : subtype p) (b : β), h a.val b) = v.prod h

theorem finsupp.​sum_subtype_domain_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} {p : α → Prop} [has_zero β] [add_comm_monoid γ] {v : α →₀ β} {h : α → β → γ} :
(∀ (x : α), x v.supportp x)(finsupp.subtype_domain p v).sum (λ (a : subtype p) (b : β), h a.val b) = v.sum h

@[simp]
theorem finsupp.​subtype_domain_add {α : Type u_1} {β : Type u_2} {p : α → Prop} [add_monoid β] {v v' : α →₀ β} :

@[instance]
def finsupp.​subtype_domain.​is_add_monoid_hom {α : Type u_1} {β : Type u_2} {p : α → Prop} [add_monoid β] :

Equations
@[simp]
theorem finsupp.​filter_add {α : Type u_1} {β : Type u_2} {p : α → Prop} [add_monoid β] {v v' : α →₀ β} :

@[instance]
def finsupp.​filter.​is_add_monoid_hom {α : Type u_1} {β : Type u_2} [add_monoid β] (p : α → Prop) :

Equations
  • _ = _
theorem finsupp.​subtype_domain_sum {α : Type u_1} {β : Type u_2} {γ : Type u_3} {p : α → Prop} [add_comm_monoid β] {s : finset γ} {h : γ → α →₀ β} :
finsupp.subtype_domain p (s.sum (λ (c : γ), h c)) = s.sum (λ (c : γ), finsupp.subtype_domain p (h c))

theorem finsupp.​subtype_domain_finsupp_sum {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [has_zero δ] {p : α → Prop} [add_comm_monoid β] {s : γ →₀ δ} {h : γ → δ → α →₀ β} :
finsupp.subtype_domain p (s.sum h) = s.sum (λ (c : γ) (d : δ), finsupp.subtype_domain p (h c d))

theorem finsupp.​filter_sum {α : Type u_1} {β : Type u_2} {γ : Type u_3} {p : α → Prop} [add_comm_monoid β] (s : finset γ) (f : γ → α →₀ β) :
finsupp.filter p (s.sum (λ (a : γ), f a)) = s.sum (λ (a : γ), finsupp.filter p (f a))

@[simp]
theorem finsupp.​subtype_domain_neg {α : Type u_1} {β : Type u_2} {p : α → Prop} [add_group β] {v : α →₀ β} :

@[simp]
theorem finsupp.​subtype_domain_sub {α : Type u_1} {β : Type u_2} {p : α → Prop} [add_group β] {v v' : α →₀ β} :

Declarations relating finsupp to multiset

def finsupp.​to_multiset {α : Type u_1} :
→₀ )multiset α

Given f : α →₀ ℕ, f.to_multiset is the multiset with multiplicities given by the values of f on the elements of α.

Equations
@[simp]
theorem finsupp.​to_multiset_zero {α : Type u_1} :

@[simp]
theorem finsupp.​to_multiset_add {α : Type u_1} (m n : α →₀ ) :

theorem finsupp.​to_multiset_single {α : Type u_1} (a : α) (n : ) :

theorem finsupp.​card_to_multiset {α : Type u_1} (f : α →₀ ) :
f.to_multiset.card = f.sum (λ (a : α), id)

theorem finsupp.​to_multiset_map {α : Type u_1} {β : Type u_2} (f : α →₀ ) (g : α → β) :

theorem finsupp.​prod_to_multiset {α : Type u_1} [comm_monoid α] (f : α →₀ ) :
f.to_multiset.prod = f.prod (λ (a : α) (n : ), a ^ n)

@[simp]
theorem finsupp.​count_to_multiset {α : Type u_1} (f : α →₀ ) (a : α) :

def finsupp.​of_multiset {α : Type u_1} :
multiset αα →₀

Given m : multiset α, of_multiset m is the finitely supported function from α to given by the multiplicities of the elements of α in m.

Equations
@[simp]
theorem finsupp.​of_multiset_apply {α : Type u_1} (m : multiset α) (a : α) :

def finsupp.​equiv_multiset {α : Type u_1} :

equiv_multiset defines an equiv between finitely supported functions from α to and multisets on α.

Equations
theorem finsupp.​mem_support_multiset_sum {α : Type u_1} {β : Type u_2} [add_comm_monoid β] {s : multiset →₀ β)} (a : α) :
a s.sum.support(∃ (f : α →₀ β) (H : f s), a f.support)

theorem finsupp.​mem_support_finset_sum {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid β] {s : finset γ} {h : γ → α →₀ β} (a : α) :
a (s.sum (λ (c : γ), h c)).support(∃ (c : γ) (H : c s), a (h c).support)

theorem finsupp.​mem_support_single {α : Type u_1} {β : Type u_2} [has_zero β] (a a' : α) (b : β) :
a (finsupp.single a' b).support a = a' b 0

Declarations about curry and uncurry

def finsupp.​curry {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid γ] :
× β →₀ γ)α →₀ β →₀ γ

Given a finitely supported function f from a product type α × β to γ, curry f is the "curried" finitely supported function from α to the type of finitely supported functions from β to γ.

Equations
theorem finsupp.​sum_curry_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [add_comm_monoid γ] [add_comm_monoid δ] (f : α × β →₀ γ) (g : α → β → γ → δ) :
(∀ (a : α) (b : β), g a b 0 = 0)(∀ (a : α) (b : β) (c₀ c₁ : γ), g a b (c₀ + c₁) = g a b c₀ + g a b c₁)f.curry.sum (λ (a : α) (f : β →₀ γ), f.sum (g a)) = f.sum (λ (p : α × β) (c : γ), g p.fst p.snd c)

def finsupp.​uncurry {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid γ] :
→₀ β →₀ γ)α × β →₀ γ

Given a finitely supported function f from α to the type of finitely supported functions from β to γ, uncurry f is the "uncurried" finitely supported function from α × β to γ.

Equations
def finsupp.​finsupp_prod_equiv {α : Type u_1} {β : Type u_2} {γ : Type u_3} [add_comm_monoid γ] :
× β →₀ γ) →₀ β →₀ γ)

finsupp_prod_equiv defines the equiv between ((α × β) →₀ γ) and (α →₀ (β →₀ γ)) given by currying and uncurrying.

Equations
theorem finsupp.​filter_curry {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [add_comm_monoid β] (f : α₁ × α₂ →₀ β) (p : α₁ → Prop) :
(finsupp.filter (λ (a : α₁ × α₂), p a.fst) f).curry = finsupp.filter p f.curry

theorem finsupp.​support_curry {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [add_comm_monoid β] (f : α₁ × α₂ →₀ β) :

def finsupp.​comap_has_scalar {α : Type u_1} {β : Type u_2} {γ : Type u_3} [group γ] [mul_action γ α] [add_comm_monoid β] :
has_scalar γ →₀ β)

Scalar multiplication by a group element g, given by precomposition with the action of g⁻¹ on the domain.

Equations
def finsupp.​comap_mul_action {α : Type u_1} {β : Type u_2} {γ : Type u_3} [group γ] [mul_action γ α] [add_comm_monoid β] :
mul_action γ →₀ β)

Scalar multiplication by a group element, given by precomposition with the action of g⁻¹ on the domain, is multiplicative in g.

Equations
def finsupp.​comap_distrib_mul_action {α : Type u_1} {β : Type u_2} {γ : Type u_3} [group γ] [mul_action γ α] [add_comm_monoid β] :

Scalar multiplication by a group element, given by precomposition with the action of g⁻¹ on the domain, is additive in the second argument.

Equations
def finsupp.​comap_distrib_mul_action_self {β : Type u_2} {γ : Type u_3} [group γ] [add_comm_monoid β] :

Scalar multiplication by a group element on finitely supported functions on a group, given by precomposition with the action of g⁻¹.

Equations
@[simp]
theorem finsupp.​comap_smul_single {α : Type u_1} {β : Type u_2} {γ : Type u_3} [group γ] [mul_action γ α] [add_comm_monoid β] (g : γ) (a : α) (b : β) :

@[simp]
theorem finsupp.​comap_smul_apply {α : Type u_1} {β : Type u_2} {γ : Type u_3} [group γ] [mul_action γ α] [add_comm_monoid β] (g : γ) (f : α →₀ β) (a : α) :
(g f) a = f (g⁻¹ a)

@[instance]
def finsupp.​has_scalar {α : Type u_1} {β : Type u_2} {γ : Type u_3} [semiring γ] [add_comm_monoid β] [semimodule γ β] :
has_scalar γ →₀ β)

Equations
@[simp]
theorem finsupp.​smul_apply' (α : Type u_1) (β : Type u_2) {γ : Type u_3} {R : semiring γ} [add_comm_monoid β] [semimodule γ β] {a : α} {b : γ} {v : α →₀ β} :
(b v) a = b v a

@[instance]
def finsupp.​semimodule (α : Type u_1) (β : Type u_2) {γ : Type u_3} [semiring γ] [add_comm_monoid β] [semimodule γ β] :
semimodule γ →₀ β)

Equations
def finsupp.​leval' {α : Type u_1} {β : Type u_2} (γ : Type u_3) [semiring γ] [add_comm_monoid β] [semimodule γ β] :
α → ((α →₀ β) →ₗ[γ] β)

Evaluation at point as a linear map. This version assumes that the codomain is a semimodule over some semiring. See also leval.

Equations
@[simp]
theorem finsupp.​coe_leval' {α : Type u_1} {β : Type u_2} (γ : Type u_3) [semiring γ] [add_comm_monoid β] [semimodule γ β] (a : α) (g : α →₀ β) :
(finsupp.leval' γ a) g = g a

def finsupp.​leval {α : Type u_1} {β : Type u_2} [semiring β] :
α → ((α →₀ β) →ₗ[β] β)

Evaluation at point as a linear map. This version assumes that the codomain is a semiring.

Equations
@[simp]
theorem finsupp.​coe_leval {α : Type u_1} {β : Type u_2} [semiring β] (a : α) (g : α →₀ β) :

theorem finsupp.​support_smul {α : Type u_1} {β : Type u_2} {γ : Type u_3} {R : semiring γ} [add_comm_monoid β] [semimodule γ β] {b : γ} {g : α →₀ β} :

@[simp]
theorem finsupp.​filter_smul {α : Type u_1} {β : Type u_2} {γ : Type u_3} {p : α → Prop} {R : semiring γ} [add_comm_monoid β] [semimodule γ β] {b : γ} {v : α →₀ β} :

theorem finsupp.​map_domain_smul {α : Type u_1} {β : Type u_2} {γ : Type u_3} {α' : Type u_4} {R : semiring γ} [add_comm_monoid β] [semimodule γ β] {f : α → α'} (b : γ) (v : α →₀ β) :

@[simp]
theorem finsupp.​smul_single {α : Type u_1} {β : Type u_2} {γ : Type u_3} {R : semiring γ} [add_comm_monoid β] [semimodule γ β] (c : γ) (a : α) (b : β) :

@[simp]
theorem finsupp.​smul_single' {α : Type u_1} {γ : Type u_3} {R : semiring γ} (c : γ) (a : α) (b : γ) :

@[simp]
theorem finsupp.​smul_apply {α : Type u_1} {β : Type u_2} [semiring β] {a : α} {b : β} {v : α →₀ β} :
(b v) a = b v a

theorem finsupp.​sum_smul_index {α : Type u_1} {β : Type u_2} {γ : Type u_3} [semiring β] [add_comm_monoid γ] {g : α →₀ β} {b : β} {h : α → β → γ} :
(∀ (i : α), h i 0 = 0)(b g).sum h = g.sum (λ (i : α) (a : β), h i (b * a))

theorem finsupp.​sum_smul_index' {α : Type u_1} {β : Type u_2} {γ : Type u_3} {δ : Type u_4} [semiring β] [add_comm_monoid γ] [semimodule β γ] [add_comm_monoid δ] {g : α →₀ γ} {b : β} {h : α → γ → δ} :
(∀ (i : α), h i 0 = 0)(b g).sum h = g.sum (λ (i : α) (c : γ), h i (b c))

theorem finsupp.​sum_mul {α : Type u_1} {β : Type u_2} {γ : Type u_3} [semiring β] [semiring γ] (b : γ) (s : α →₀ β) {f : α → β → γ} :
s.sum f * b = s.sum (λ (a : α) (c : β), f a (s a) * b)

theorem finsupp.​mul_sum {α : Type u_1} {β : Type u_2} {γ : Type u_3} [semiring β] [semiring γ] (b : γ) (s : α →₀ β) {f : α → β → γ} :
b * s.sum f = s.sum (λ (a : α) (c : β), b * f a (s a))

theorem finsupp.​eq_zero_of_zero_eq_one {α : Type u_1} {β : Type u_2} [semiring β] (zero_eq_one : 0 = 1) (l : α →₀ β) :
l = 0

def finsupp.​restrict_support_equiv {α : Type u_1} (s : set α) (β : Type u_2) [add_comm_monoid β] :
{f // (f.support) s} (s →₀ β)

Given an add_comm_monoid β and s : set α, restrict_support_equiv s β is the equiv between the subtype of finitely supported functions with support contained in s and the type of finitely supported functions from s.

Equations
def finsupp.​dom_congr {β : Type u_2} {α₁ : Type u_6} {α₂ : Type u_7} [add_comm_monoid β] :
α₁ α₂(α₁ →₀ β) (α₂ →₀ β)

Given add_comm_monoid β and e : α₁ ≃ α₂, dom_congr e is the corresponding equiv between α₁ →₀ β and α₂ →₀ β.

Equations

Declarations about sigma types

def finsupp.​split {β : Type u_2} {ι : Type u_5} {αs : ι → Type u_10} [has_zero β] (l : (Σ (i : ι), αs i) →₀ β) (i : ι) :
αs i →₀ β

Given l, a finitely supported function from the sigma type Σ (i : ι), αs i to β and an index element i : ι, split l i is the ith component of l, a finitely supported function from as i to β.

Equations
theorem finsupp.​split_apply {β : Type u_2} {ι : Type u_5} {αs : ι → Type u_10} [has_zero β] (l : (Σ (i : ι), αs i) →₀ β) (i : ι) (x : αs i) :
(l.split i) x = l i, x⟩

def finsupp.​split_support {β : Type u_2} {ι : Type u_5} {αs : ι → Type u_10} [has_zero β] :
((Σ (i : ι), αs i) →₀ β)finset ι

Given l, a finitely supported function from the sigma type Σ (i : ι), αs i to β, split_support l is the finset of indices in ι that appear in the support of l.

Equations
theorem finsupp.​mem_split_support_iff_nonzero {β : Type u_2} {ι : Type u_5} {αs : ι → Type u_10} [has_zero β] (l : (Σ (i : ι), αs i) →₀ β) (i : ι) :

def finsupp.​split_comp {β : Type u_2} {γ : Type u_3} {ι : Type u_5} {αs : ι → Type u_10} [has_zero β] (l : (Σ (i : ι), αs i) →₀ β) [has_zero γ] (g : Π (i : ι), (αs i →₀ β) → γ) :
(∀ (i : ι) (x : αs i →₀ β), x = 0 g i x = 0)ι →₀ γ

Given l, a finitely supported function from the sigma type Σ i, αs i to β and an ι-indexed family g of functions from (αs i →₀ β) to γ, split_comp defines a finitely supported function from the index type ι to γ given by composing g i with split l i.

Equations
theorem finsupp.​sigma_support {β : Type u_2} {ι : Type u_5} {αs : ι → Type u_10} [has_zero β] (l : (Σ (i : ι), αs i) →₀ β) :
l.support = l.split_support.sigma (λ (i : ι), (l.split i).support)

theorem finsupp.​sigma_sum {β : Type u_2} {γ : Type u_3} {ι : Type u_5} {αs : ι → Type u_10} [has_zero β] (l : (Σ (i : ι), αs i) →₀ β) [add_comm_monoid γ] (f : (Σ (i : ι), αs i)β → γ) :
l.sum f = l.split_support.sum (λ (i : ι), (l.split i).sum (λ (a : αs i) (b : β), f i, a⟩ b))

Declarations relating multiset to finsupp

def multiset.​to_finsupp {α : Type u_1} :
multiset αα →₀

Given a multiset s, s.to_finsupp returns the finitely supported function on given by the multiplicities of the elements of s.

Equations
@[simp]
theorem multiset.​to_finsupp_support {α : Type u_1} (s : multiset α) :

@[simp]
theorem multiset.​to_finsupp_apply {α : Type u_1} (s : multiset α) (a : α) :

@[simp]
theorem multiset.​to_finsupp_zero {α : Type u_1} :

@[simp]
theorem multiset.​to_finsupp_add {α : Type u_1} (s t : multiset α) :

theorem multiset.​to_finsupp_singleton {α : Type u_1} (a : α) :

@[simp]
theorem multiset.​to_finsupp_to_multiset {α : Type u_1} (s : multiset α) :

Declarations about order(ed) instances on finsupp

@[instance]
def finsupp.​preorder {α : Type u_1} {σ : Type u_10} [preorder α] [has_zero α] :

Equations
@[instance]
def finsupp.​partial_order {α : Type u_1} {σ : Type u_10} [partial_order α] [has_zero α] :

Equations
theorem finsupp.​le_iff {α : Type u_1} {σ : Type u_10} [canonically_ordered_add_monoid α] (f g : σ →₀ α) :
f g ∀ (s : σ), s f.supportf s g s

@[simp]
theorem finsupp.​add_eq_zero_iff {α : Type u_1} {σ : Type u_10} [canonically_ordered_add_monoid α] (f g : σ →₀ α) :
f + g = 0 f = 0 g = 0

@[simp]
theorem finsupp.​to_multiset_to_finsupp {σ : Type u_10} (f : σ →₀ ) :

theorem finsupp.​sum_id_lt_of_lt {σ : Type u_10} (m n : σ →₀ ) :
m < nm.sum (λ (_x : σ), id) < n.sum (λ (_x : σ), id)

theorem finsupp.​lt_wf (σ : Type u_10) :

The order on σ →₀ ℕ is well-founded.

@[instance]

Equations
def finsupp.​antidiagonal {σ : Type u_10} :
→₀ )→₀ ) × →₀ ) →₀

The finsupp counterpart of multiset.antidiagonal: the antidiagonal of s : σ →₀ ℕ consists of all pairs (t₁, t₂) : (σ →₀ ℕ) × (σ →₀ ℕ) such that t₁ + t₂ = s. The finitely supported function antidiagonal s is equal to the multiplicities of these pairs.

Equations
theorem finsupp.​mem_antidiagonal_support {σ : Type u_10} {f : σ →₀ } {p : →₀ ) × →₀ )} :

@[simp]
theorem finsupp.​antidiagonal_zero {σ : Type u_10} :

theorem finsupp.​swap_mem_antidiagonal_support {σ : Type u_10} {n : σ →₀ } {f : →₀ ) × →₀ )} :

theorem finsupp.​finite_le_nat {σ : Type u_10} (n : σ →₀ ) :
{m : σ →₀ | m n}.finite

Let n : σ →₀ ℕ be a finitely supported function. The set of m : σ →₀ ℕ that are coordinatewise less than or equal to n, is a finite set.

theorem finsupp.​finite_lt_nat {σ : Type u_10} (n : σ →₀ ) :
{m : σ →₀ | m < n}.finite

Let n : σ →₀ ℕ be a finitely supported function. The set of m : σ →₀ ℕ that are coordinatewise less than or equal to n, but not equal to n everywhere, is a finite set.