@@ -10,8 +10,8 @@ import Mathlib.Algebra.Support
10
10
/-!
11
11
# Indicator function
12
12
13
- - `indicator (s : Set α) (f : α → β) (a : α)` is `f a` if `a ∈ s` and is `0` otherwise.
14
- - `mulIndicator (s : Set α) (f : α → β) (a : α)` is `f a` if `a ∈ s` and is `1` otherwise.
13
+ - `Set. indicator (s : Set α) (f : α → β) (a : α)` is `f a` if `a ∈ s` and is `0` otherwise.
14
+ - `Set. mulIndicator (s : Set α) (f : α → β) (a : α)` is `f a` if `a ∈ s` and is `1` otherwise.
15
15
16
16
17
17
## Implementation note
@@ -21,10 +21,10 @@ used to indicate membership of an element in a set `s`,
21
21
having the value `1` for all elements of `s` and the value `0` otherwise.
22
22
But since it is usually used to restrict a function to a certain set `s`,
23
23
we let the indicator function take the value `f x` for some function `f`, instead of `1`.
24
- If the usual indicator function is needed, just set `f` to be the constant function `λx, 1`.
24
+ If the usual indicator function is needed, just set `f` to be the constant function `fun _ ↦ 1`.
25
25
26
26
The indicator function is implemented non-computably, to avoid having to pass around `Decidable`
27
- arguments. This is in contrast with the design of `Pi.Single ` or `Set.Piecewise `.
27
+ arguments. This is in contrast with the design of `Pi.single ` or `Set.piecewise `.
28
28
29
29
## Tags
30
30
indicator, characteristic
@@ -34,27 +34,19 @@ open BigOperators
34
34
35
35
open Function
36
36
37
- variable {α β ι M N : Type _ }
37
+ variable {α β ι M N : Type * }
38
38
39
39
namespace Set
40
40
41
41
section One
42
42
43
43
variable [One M] [One N] {s t : Set α} {f g : α → M} {a : α}
44
44
45
- /-- `indicator s f a` is `f a` if `a ∈ s`, `0` otherwise. -/
46
- noncomputable def indicator {M} [Zero M] (s : Set α) (f : α → M) : α → M
47
- | x =>
48
- haveI := Classical.decPred (· ∈ s)
49
- if x ∈ s then f x else 0
50
- #align set.indicator Set.indicator
51
-
52
- /-- `mulIndicator s f a` is `f a` if `a ∈ s`, `1` otherwise. -/
53
- @[to_additive existing]
54
- noncomputable def mulIndicator (s : Set α) (f : α → M) : α → M
55
- | x =>
56
- haveI := Classical.decPred (· ∈ s)
57
- if x ∈ s then f x else 1
45
+ /-- `Set.mulIndicator s f a` is `f a` if `a ∈ s`, `1` otherwise. -/
46
+ @[to_additive "`Set.indicator s f a` is `f a` if `a ∈ s`, `0` otherwise."]
47
+ noncomputable def mulIndicator (s : Set α) (f : α → M) (x : α) : M :=
48
+ haveI := Classical.decPred (· ∈ s)
49
+ if x ∈ s then f x else 1
58
50
#align set.mul_indicator Set.mulIndicator
59
51
60
52
@[to_additive (attr := simp)]
@@ -67,21 +59,19 @@ theorem piecewise_eq_mulIndicator [DecidablePred (· ∈ s)] : s.piecewise f 1 =
67
59
@[to_additive]
68
60
theorem mulIndicator_apply (s : Set α) (f : α → M) (a : α) [Decidable (a ∈ s)] :
69
61
mulIndicator s f a = if a ∈ s then f a else 1 := by
70
- unfold mulIndicator
71
- split_ifs with h <;> simp
62
+ unfold mulIndicator
63
+ congr
72
64
#align set.mul_indicator_apply Set.mulIndicator_apply
73
65
#align set.indicator_apply Set.indicator_apply
74
66
75
67
@[to_additive (attr := simp)]
76
68
theorem mulIndicator_of_mem (h : a ∈ s) (f : α → M) : mulIndicator s f a = f a :=
77
- letI := Classical.dec (a ∈ s)
78
69
if_pos h
79
70
#align set.mul_indicator_of_mem Set.mulIndicator_of_mem
80
71
#align set.indicator_of_mem Set.indicator_of_mem
81
72
82
73
@[to_additive (attr := simp)]
83
74
theorem mulIndicator_of_not_mem (h : a ∉ s) (f : α → M) : mulIndicator s f a = 1 :=
84
- letI := Classical.dec (a ∈ s)
85
75
if_neg h
86
76
#align set.mul_indicator_of_not_mem Set.mulIndicator_of_not_mem
87
77
#align set.indicator_of_not_mem Set.indicator_of_not_mem
@@ -233,8 +223,7 @@ theorem mulIndicator_mulIndicator (s t : Set α) (f : α → M) :
233
223
mulIndicator s (mulIndicator t f) = mulIndicator (s ∩ t) f :=
234
224
funext fun x => by
235
225
simp only [mulIndicator]
236
- split_ifs
237
- repeat' simp_all (config := { contextual := true })
226
+ split_ifs <;> simp_all (config := { contextual := true })
238
227
#align set.mul_indicator_mul_indicator Set.mulIndicator_mulIndicator
239
228
#align set.indicator_indicator Set.indicator_indicator
240
229
@@ -570,8 +559,9 @@ theorem mulIndicator_compl (s : Set α) (f : α → G) :
570
559
#align set.mul_indicator_compl Set.mulIndicator_compl
571
560
#align set.indicator_compl' Set.indicator_compl'
572
561
573
- theorem indicator_compl {G} [AddGroup G] (s : Set α) (f : α → G) :
574
- indicator sᶜ f = f - indicator s f := by rw [sub_eq_add_neg, indicator_compl']
562
+ @[to_additive indicator_compl]
563
+ theorem mulIndicator_compl' (s : Set α) (f : α → G) :
564
+ mulIndicator sᶜ f = f / mulIndicator s f := by rw [div_eq_mul_inv, mulIndicator_compl]
575
565
#align set.indicator_compl Set.indicator_compl
576
566
577
567
@[to_additive indicator_diff']
@@ -584,8 +574,10 @@ theorem mulIndicator_diff (h : s ⊆ t) (f : α → G) :
584
574
#align set.mul_indicator_diff Set.mulIndicator_diff
585
575
#align set.indicator_diff' Set.indicator_diff'
586
576
587
- theorem indicator_diff {G : Type *} [AddGroup G] {s t : Set α} (h : s ⊆ t) (f : α → G) :
588
- indicator (t \ s) f = indicator t f - indicator s f := by rw [indicator_diff' h, sub_eq_add_neg]
577
+ @[to_additive indicator_diff]
578
+ theorem mulIndicator_diff' (h : s ⊆ t) (f : α → G) :
579
+ mulIndicator (t \ s) f = mulIndicator t f / mulIndicator s f := by
580
+ rw [mulIndicator_diff h, div_eq_mul_inv]
589
581
#align set.indicator_diff Set.indicator_diff
590
582
591
583
end Group
@@ -658,7 +650,7 @@ theorem mulIndicator_finset_prod (I : Finset ι) (s : Set α) (f : ι → α →
658
650
#align set.indicator_finset_sum Set.indicator_finset_sum
659
651
660
652
@[to_additive]
661
- theorem mulIndicator_finset_biUnion {ι} (I : Finset ι) (s : ι → Set α) {f : α → M} :
653
+ theorem mulIndicator_finset_biUnion (I : Finset ι) (s : ι → Set α) {f : α → M} :
662
654
(∀ i ∈ I, ∀ j ∈ I, i ≠ j → Disjoint (s i) (s j)) →
663
655
mulIndicator (⋃ i ∈ I, s i) f = fun a => ∏ i in I, mulIndicator (s i) f a := by
664
656
classical
@@ -681,7 +673,7 @@ theorem mulIndicator_finset_biUnion {ι} (I : Finset ι) (s : ι → Set α) {f
681
673
#align set.indicator_finset_bUnion Set.indicator_finset_biUnion
682
674
683
675
@[to_additive]
684
- theorem mulIndicator_finset_biUnion_apply {ι} (I : Finset ι) (s : ι → Set α) {f : α → M}
676
+ theorem mulIndicator_finset_biUnion_apply (I : Finset ι) (s : ι → Set α) {f : α → M}
685
677
(h : ∀ i ∈ I, ∀ j ∈ I, i ≠ j → Disjoint (s i) (s j)) (x : α) :
686
678
mulIndicator (⋃ i ∈ I, s i) f x = ∏ i in I, mulIndicator (s i) f x := by
687
679
rw [Set.mulIndicator_finset_biUnion I s h]
@@ -854,8 +846,8 @@ theorem mulIndicator_le_self' (hf : ∀ (x) (_ : x ∉ s), 1 ≤ f x) : mulIndic
854
846
#align set.indicator_le_self' Set.indicator_le_self'
855
847
856
848
@[to_additive]
857
- theorem mulIndicator_iUnion_apply {ι M } [CompleteLattice M] [One M] (h1 : (⊥ : M) = 1 )
858
- (s : ι → Set α) (f : α → M) (x : α) :
849
+ theorem mulIndicator_iUnion_apply {ι : Sort *} {M : Type * } [CompleteLattice M] [One M]
850
+ (h1 : (⊥ : M) = 1 ) ( s : ι → Set α) (f : α → M) (x : α) :
859
851
mulIndicator (⋃ i, s i) f x = ⨆ i, mulIndicator (s i) f x := by
860
852
by_cases hx : x ∈ ⋃ i, s i
861
853
· rw [mulIndicator_of_mem hx]
@@ -869,8 +861,8 @@ theorem mulIndicator_iUnion_apply {ι M} [CompleteLattice M] [One M] (h1 : (⊥
869
861
#align set.mul_indicator_Union_apply Set.mulIndicator_iUnion_apply
870
862
#align set.indicator_Union_apply Set.indicator_iUnion_apply
871
863
872
- @[to_additive] lemma mulIndicator_iInter_apply {ι M} [Nonempty ι] [CompleteLattice M] [One M ]
873
- (h1 : (⊥ : M) = 1 ) (s : ι → Set α) (f : α → M) (x : α) :
864
+ @[to_additive] lemma mulIndicator_iInter_apply {ι : Sort *} {M : Type *} [Nonempty ι ]
865
+ [CompleteLattice M] [One M] (h1 : (⊥ : M) = 1 ) (s : ι → Set α) (f : α → M) (x : α) :
874
866
mulIndicator (⋂ i, s i) f x = ⨅ i, mulIndicator (s i) f x := by
875
867
by_cases hx : x ∈ ⋂ i, s i
876
868
· rw [mulIndicator_of_mem hx]
0 commit comments