Support of a function #
In this file we define Function.support f = {x | f x ≠ 0}
and prove its basic properties.
We also define Function.mulSupport f = {x | f x ≠ 1}
.
mulSupport
of a function is the set of points x
such that f x ≠ 1
.
Instances For
@[simp]
theorem
Function.mulSupport_subset_iff'
{α : Type u_1}
{M : Type u_5}
[One M]
{f : α → M}
{s : Set α}
:
theorem
Function.mulSupport_update_of_ne_one
{α : Type u_1}
{M : Type u_5}
[One M]
[DecidableEq α]
(f : α → M)
(x : α)
{y : M}
(hy : y ≠ 1)
:
theorem
Function.mulSupport_update_one
{α : Type u_1}
{M : Type u_5}
[One M]
[DecidableEq α]
(f : α → M)
(x : α)
:
theorem
Function.mulSupport_update_eq_ite
{α : Type u_1}
{M : Type u_5}
[One M]
[DecidableEq α]
[DecidableEq M]
(f : α → M)
(x : α)
(y : M)
:
theorem
Function.mulSupport_extend_one_subset
{α : Type u_1}
{M' : Type u_6}
{N : Type u_7}
[One N]
{f : α → M'}
{g : α → N}
:
theorem
Function.mulSupport_disjoint_iff
{α : Type u_1}
{M : Type u_5}
[One M]
{f : α → M}
{s : Set α}
:
theorem
Function.disjoint_mulSupport_iff
{α : Type u_1}
{M : Type u_5}
[One M]
{f : α → M}
{s : Set α}
:
@[simp]
@[simp]
theorem
Function.range_subset_insert_image_mulSupport
{α : Type u_1}
{M : Type u_5}
[One M]
(f : α → M)
:
@[simp]
@[simp]
theorem
Function.mulSupport_comp_eq_preimage
{α : Type u_1}
{β : Type u_2}
{M : Type u_5}
[One M]
(g : β → M)
(f : α → β)
:
@[simp]
@[simp]
@[simp]
@[simp]
theorem
Function.mulSupport_mul_inv
{α : Type u_1}
{G : Type u_9}
[DivisionMonoid G]
(f g : α → G)
:
theorem
Pi.mulSupport_mulSingle_subset
{A : Type u_1}
{B : Type u_2}
[DecidableEq A]
[One B]
{a : A}
{b : B}
:
theorem
Pi.support_single_subset
{A : Type u_1}
{B : Type u_2}
[DecidableEq A]
[Zero B]
{a : A}
{b : B}
:
@[simp]
theorem
Pi.mulSupport_mulSingle_of_ne
{A : Type u_1}
{B : Type u_2}
[DecidableEq A]
[One B]
{a : A}
{b : B}
(h : b ≠ 1)
:
@[simp]
theorem
Pi.support_single_of_ne
{A : Type u_1}
{B : Type u_2}
[DecidableEq A]
[Zero B]
{a : A}
{b : B}
(h : b ≠ 0)
:
theorem
Pi.mulSupport_mulSingle
{A : Type u_1}
{B : Type u_2}
[DecidableEq A]
[One B]
{a : A}
{b : B}
[DecidableEq B]
:
theorem
Pi.support_single
{A : Type u_1}
{B : Type u_2}
[DecidableEq A]
[Zero B]
{a : A}
{b : B}
[DecidableEq B]
:
theorem
Pi.mulSupport_mulSingle_disjoint
{A : Type u_1}
{B : Type u_2}
[DecidableEq A]
[One B]
{b b' : B}
(hb : b ≠ 1)
(hb' : b' ≠ 1)
{i j : A}
:
theorem
Pi.support_single_disjoint
{A : Type u_1}
{B : Type u_2}
[DecidableEq A]
[Zero B]
{b b' : B}
(hb : b ≠ 0)
(hb' : b' ≠ 0)
{i j : A}
: