Skip to content

Commit

Permalink
feat(topology/locally_constant): Characteristic functions on clopen s…
Browse files Browse the repository at this point in the history
…ets are locally constant (#11708)

Gives an API for characteristic functions on clopen sets, `char_fn`, which are locally constant functions.



Co-authored-by: Kevin Buzzard <k.buzzard@imperial.ac.uk>
Co-authored-by: Floris van Doorn <fpvdoorn@gmail.com>
  • Loading branch information
3 people committed May 19, 2022
1 parent d403cad commit ffe7002
Show file tree
Hide file tree
Showing 3 changed files with 79 additions and 0 deletions.
13 changes: 13 additions & 0 deletions src/algebra/indicator_function.lean
Expand Up @@ -459,6 +459,19 @@ begin
simp [indicator_apply, ← ite_and],
end

variables (M) [nontrivial M]

lemma indicator_eq_zero_iff_not_mem {U : set α} {x : α} :
indicator U 1 x = (0 : M) ↔ x ∉ U :=
by { classical, simp [indicator_apply, imp_false] }

lemma indicator_eq_one_iff_mem {U : set α} {x : α} :
indicator U 1 x = (1 : M) ↔ x ∈ U :=
by { classical, simp [indicator_apply, imp_false] }

lemma indicator_one_inj {U V : set α} (h : indicator U (1 : α → M) = indicator V 1) : U = V :=
by { ext, simp_rw [← indicator_eq_one_iff_mem M, h] }

end mul_zero_one_class

section order
Expand Down
23 changes: 23 additions & 0 deletions src/topology/locally_constant/algebra.lean
Expand Up @@ -72,6 +72,29 @@ instance [mul_zero_class Y] : mul_zero_class (locally_constant X Y) :=
instance [mul_zero_one_class Y] : mul_zero_one_class (locally_constant X Y) :=
{ .. locally_constant.mul_zero_class, .. locally_constant.mul_one_class }

section char_fn

variables (Y) [mul_zero_one_class Y] {U V : set X}

/-- Characteristic functions are locally constant functions taking `x : X` to `1` if `x ∈ U`,
where `U` is a clopen set, and `0` otherwise. -/
noncomputable def char_fn (hU : is_clopen U) : locally_constant X Y := indicator 1 hU

lemma coe_char_fn (hU : is_clopen U) : (char_fn Y hU : X → Y) = set.indicator U 1 :=
rfl

lemma char_fn_eq_one [nontrivial Y] (x : X) (hU : is_clopen U) :
char_fn Y hU x = (1 : Y) ↔ x ∈ U := set.indicator_eq_one_iff_mem _

lemma char_fn_eq_zero [nontrivial Y] (x : X) (hU : is_clopen U) :
char_fn Y hU x = (0 : Y) ↔ x ∉ U := set.indicator_eq_zero_iff_not_mem _

lemma char_fn_inj [nontrivial Y] (hU : is_clopen U) (hV : is_clopen V)
(h : char_fn Y hU = char_fn Y hV) : U = V :=
set.indicator_one_inj Y $ coe_inj.mpr h

end char_fn

@[to_additive] instance [has_div Y] : has_div (locally_constant X Y) :=
{ div := λ f g, ⟨f / g, f.is_locally_constant.div g.is_locally_constant⟩ }

Expand Down
43 changes: 43 additions & 0 deletions src/topology/locally_constant/basic.lean
Expand Up @@ -426,4 +426,47 @@ lemma coe_desc {X α β : Type*} [topological_space X] (f : X → α) (g : α

end desc

section indicator
variables {R : Type*} [has_one R] {U : set X} (f : locally_constant X R)
open_locale classical

/-- Given a clopen set `U` and a locally constant function `f`, `locally_constant.mul_indicator`
returns the locally constant function that is `f` on `U` and `1` otherwise. -/
@[to_additive /-" Given a clopen set `U` and a locally constant function `f`,
`locally_constant.indicator` returns the locally constant function that is `f` on `U` and `0`
otherwise. "-/, simps]
noncomputable def mul_indicator (hU : is_clopen U) :
locally_constant X R :=
{ to_fun := set.mul_indicator U f,
is_locally_constant :=
begin
rw is_locally_constant.iff_exists_open, rintros x,
obtain ⟨V, hV, hx, h'⟩ := (is_locally_constant.iff_exists_open _).1 f.is_locally_constant x,
by_cases x ∈ U,
{ refine ⟨U ∩ V, is_open.inter hU.1 hV, set.mem_inter h hx, _⟩, rintros y hy,
rw set.mem_inter_iff at hy, rw [set.mul_indicator_of_mem hy.1, set.mul_indicator_of_mem h],
apply h' y hy.2, },
{ rw ←set.mem_compl_iff at h, refine ⟨Uᶜ, (is_clopen.compl hU).1, h, _⟩,
rintros y hy, rw set.mem_compl_iff at h, rw set.mem_compl_iff at hy,
simp [h, hy], },
end, }

variables (a : X)

@[to_additive]
theorem mul_indicator_apply_eq_if (hU : is_clopen U) :
mul_indicator f hU a = if a ∈ U then f a else 1 :=
set.mul_indicator_apply U f a

variables {a}

@[to_additive]
theorem mul_indicator_of_mem (hU : is_clopen U) (h : a ∈ U) : f.mul_indicator hU a = f a :=
by{ rw mul_indicator_apply, apply set.mul_indicator_of_mem h, }

@[to_additive]
theorem mul_indicator_of_not_mem (hU : is_clopen U) (h : a ∉ U) : f.mul_indicator hU a = 1 :=
by{ rw mul_indicator_apply, apply set.mul_indicator_of_not_mem h, }

end indicator
end locally_constant

0 comments on commit ffe7002

Please sign in to comment.