Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 10ebb71

Browse files
fpvandoornsgouezel
andcommitted
feat(measure_theory): induction principles in measure theory (#3978)
This commit adds three induction principles for measure theory * To prove something for arbitrary simple functions * To prove something for arbitrary measurable functions into `ennreal` * To prove something for arbitrary measurable + integrable functions. This also adds some basic lemmas to various files. Not all of them are used in this PR, some will be used by near-future PRs. Co-authored-by: sgouezel <sebastien.gouezel@univ-rennes1.fr>
1 parent bf7487b commit 10ebb71

File tree

11 files changed

+250
-26
lines changed

11 files changed

+250
-26
lines changed

src/algebra/group/pi.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -79,6 +79,19 @@ instance ordered_comm_group [∀ i, ordered_comm_group $ f i] :
7979
..pi.comm_group,
8080
..pi.partial_order }
8181

82+
section instance_lemmas
83+
open function
84+
85+
variables {α β γ : Type*}
86+
87+
@[simp, to_additive] lemma const_one [has_one β] : const α (1 : β) = 1 := rfl
88+
89+
@[simp, to_additive] lemma comp_one [has_one β] {f : β → γ} : f ∘ 1 = const α (f 1) := rfl
90+
91+
@[simp, to_additive] lemma one_comp [has_one γ] {f : α → β} : (1 : β → γ) ∘ f = 1 := rfl
92+
93+
end instance_lemmas
94+
8295
variables [decidable_eq I]
8396
variables [Π i, has_zero (f i)]
8497

src/data/indicator_function.lean

Lines changed: 32 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -27,9 +27,9 @@ indicator, characteristic
2727

2828
noncomputable theory
2929
open_locale classical big_operators
30+
open function
3031

31-
universes u v
32-
variables {α : Type u} {β : Type v}
32+
variables {α β γ : Type*}
3333

3434
namespace set
3535

@@ -40,6 +40,8 @@ variables [has_zero β] {s t : set α} {f g : α → β} {a : α}
4040
@[reducible]
4141
def indicator (s : set α) (f : α → β) : α → β := λ x, if x ∈ s then f x else 0
4242

43+
@[simp] lemma piecewise_eq_indicator {s : set α} : s.piecewise f 0 = s.indicator f := rfl
44+
4345
lemma indicator_apply (s : set α) (f : α → β) (a : α) :
4446
indicator s f a = if a ∈ s then f a else 0 := rfl
4547

@@ -71,14 +73,23 @@ funext $ λx, indicator_of_mem (mem_univ _) f
7173
funext $ λx, indicator_of_not_mem (not_mem_empty _) f
7274

7375
variable (β)
76+
7477
@[simp] lemma indicator_zero (s : set α) : indicator s (λx, (0:β)) = λx, (0:β) :=
7578
funext $ λx, by { simp only [indicator], split_ifs, refl, refl }
79+
80+
@[simp] lemma indicator_zero' {s : set α} : s.indicator (0 : α → β) = 0 :=
81+
indicator_zero β s
82+
7683
variable {β}
7784

7885
lemma indicator_indicator (s t : set α) (f : α → β) : indicator s (indicator t f) = indicator (s ∩ t) f :=
7986
funext $ λx, by { simp only [indicator], split_ifs, repeat {simp * at * {contextual := tt}} }
8087

81-
lemma indicator_comp_of_zero {γ} [has_zero γ] {g : β → γ} (hg : g 0 = 0) :
88+
lemma comp_indicator (h : β → γ) (f : α → β) {s : set α} {x : α} :
89+
h (s.indicator f x) = s.piecewise (h ∘ f) (const α (h 0)) x :=
90+
s.comp_piecewise h
91+
92+
lemma indicator_comp_of_zero [has_zero γ] {g : β → γ} (hg : g 0 = 0) :
8293
indicator s (g ∘ f) = g ∘ (indicator s f) :=
8394
begin
8495
funext,
@@ -178,6 +189,24 @@ lemma indicator_smul (s : set α) (r : 𝕜) (f : α → β) :
178189
indicator s (λ (x : α), r • f x) = λ (x : α), r • indicator s f x :=
179190
by { simp only [indicator], funext, split_ifs, refl, exact (smul_zero r).symm }
180191

192+
lemma indicator_add_eq_left {f g : α → β} (h : univ ⊆ f ⁻¹' {0} ∪ g ⁻¹' {0}) :
193+
(f ⁻¹' {0})ᶜ.indicator (f + g) = f :=
194+
begin
195+
ext x, by_cases hx : x ∈ (f ⁻¹' {0})ᶜ,
196+
{ have : g x = 0, { simp at hx, specialize h (mem_univ x), simpa [hx] using h },
197+
simp [hx, this] },
198+
{ simp * at * }
199+
end
200+
201+
lemma indicator_add_eq_right {f g : α → β} (h : univ ⊆ f ⁻¹' {0} ∪ g ⁻¹' {0}) :
202+
(g ⁻¹' {0})ᶜ.indicator (f + g) = g :=
203+
begin
204+
ext x, by_cases hx : x ∈ (g ⁻¹' {0})ᶜ,
205+
{ have : f x = 0, { simp at hx, specialize h (mem_univ x), simpa [hx] using h },
206+
simp [hx, this] },
207+
{ simp * at * }
208+
end
209+
181210
end add_monoid
182211

183212
section add_group

src/data/list/indexes.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -147,7 +147,7 @@ theorem mmap_with_index'_aux_eq_mmap_with_index_aux {α} (f : ℕ → α → m p
147147
mmap_with_index'_aux f start as =
148148
mmap_with_index_aux f start as *> pure punit.star :=
149149
by induction as generalizing start;
150-
simp [mmap_with_index'_aux, mmap_with_index_aux, *, seq_right_eq, const]
150+
simp [mmap_with_index'_aux, mmap_with_index_aux, *, seq_right_eq, const, -comp_const]
151151
with functor_norm
152152

153153
theorem mmap_with_index'_eq_mmap_with_index {α} (f : ℕ → α → m punit) (as : list α) :

src/data/real/ennreal.lean

Lines changed: 22 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -182,6 +182,8 @@ end
182182
protected lemma lt_top_iff_ne_top : a < ∞ ↔ a ≠ ∞ := lt_top_iff_ne_top
183183
protected lemma bot_lt_iff_ne_bot : 0 < a ↔ a ≠ 0 := bot_lt_iff_ne_bot
184184

185+
lemma not_lt_top {x : ennreal} : ¬ x < ∞ ↔ x = ∞ := by rw [lt_top_iff_ne_top, not_not]
186+
185187
lemma add_ne_top : a + b ≠ ∞ ↔ a ≠ ∞ ∧ b ≠ ∞ :=
186188
by simpa only [lt_top_iff_ne_top] using add_lt_top
187189

@@ -197,16 +199,28 @@ lemma top_pow {n:ℕ} (h : 0 < n) : ∞^n = ∞ :=
197199
nat.le_induction (pow_one _) (λ m hm hm', by rw [pow_succ, hm', top_mul_top])
198200
_ (nat.succ_le_of_lt h)
199201

200-
lemma mul_eq_top {a b : ennreal} : a * b = ⊤ ↔ (a ≠ 0 ∧ b = ⊤) ∨ (a = ⊤ ∧ b ≠ 0) :=
202+
lemma mul_eq_top : a * b = ⊤ ↔ (a ≠ 0 ∧ b = ⊤) ∨ (a = ⊤ ∧ b ≠ 0) :=
201203
with_top.mul_eq_top_iff
202204

203-
lemma mul_ne_top {a b : ennreal} : a ≠ ∞ → b ≠ ∞ → a * b ≠ ∞ :=
205+
lemma mul_ne_top : a ≠ ∞ → b ≠ ∞ → a * b ≠ ∞ :=
204206
by simp [(≠), mul_eq_top] {contextual := tt}
205207

206-
lemma mul_lt_top {a b : ennreal} : a < ⊤ → b < ⊤ → a * b < ⊤ :=
208+
lemma mul_lt_top : a < ⊤ → b < ⊤ → a * b < ⊤ :=
207209
by simpa only [ennreal.lt_top_iff_ne_top] using mul_ne_top
208210

209-
@[simp] lemma mul_pos {a b : ennreal} : 0 < a * b ↔ 0 < a ∧ 0 < b :=
211+
lemma ne_top_of_mul_ne_top_left (h : a * b ≠ ⊤) (hb : b ≠ 0) : a ≠ ⊤ :=
212+
by { simp [mul_eq_top, hb, not_or_distrib] at h ⊢, exact h.2 }
213+
214+
lemma ne_top_of_mul_ne_top_right (h : a * b ≠ ⊤) (ha : a ≠ 0) : b ≠ ⊤ :=
215+
ne_top_of_mul_ne_top_left (by rwa [mul_comm]) ha
216+
217+
lemma lt_top_of_mul_lt_top_left (h : a * b < ⊤) (hb : b ≠ 0) : a < ⊤ :=
218+
by { rw [ennreal.lt_top_iff_ne_top] at h ⊢, exact ne_top_of_mul_ne_top_left h hb }
219+
220+
lemma lt_top_of_mul_lt_top_right (h : a * b < ⊤) (ha : a ≠ 0) : b < ⊤ :=
221+
lt_top_of_mul_lt_top_left (by rwa [mul_comm]) ha
222+
223+
@[simp] lemma mul_pos : 0 < a * b ↔ 0 < a ∧ 0 < b :=
210224
by simp only [zero_lt_iff_ne_zero, ne.def, mul_eq_zero, not_or_distrib]
211225

212226
lemma pow_eq_top : ∀ n:ℕ, a^n=∞ → a=∞
@@ -936,10 +950,10 @@ by rw [← two_mul, ← div_def, div_self two_ne_zero two_ne_top]
936950
lemma add_halves (a : ennreal) : a / 2 + a / 2 = a :=
937951
by rw [div_def, ← mul_add, inv_two_add_inv_two, mul_one]
938952

939-
@[simp] lemma div_zero_iff {a b : ennreal} : a / b = 0 ↔ a = 0 ∨ b = ⊤ :=
953+
@[simp] lemma div_zero_iff : a / b = 0 ↔ a = 0 ∨ b = ⊤ :=
940954
by simp [div_def, mul_eq_zero]
941955

942-
@[simp] lemma div_pos_iff {a b : ennreal} : 0 < a / b ↔ a ≠ 0 ∧ b ≠ ⊤ :=
956+
@[simp] lemma div_pos_iff : 0 < a / b ↔ a ≠ 0 ∧ b ≠ ⊤ :=
943957
by simp [zero_lt_iff_ne_zero, not_or_distrib]
944958

945959
lemma half_pos {a : ennreal} (h : 0 < a) : 0 < a / 2 :=
@@ -1113,7 +1127,7 @@ end
11131127
@[simp] lemma to_real_top_mul (a : ennreal) : ennreal.to_real (⊤ * a) = 0 :=
11141128
by { rw mul_comm, exact to_real_mul_top _ }
11151129

1116-
lemma to_real_eq_to_real {a b : ennreal} (ha : a < ⊤) (hb : b < ⊤) :
1130+
lemma to_real_eq_to_real (ha : a < ⊤) (hb : b < ⊤) :
11171131
ennreal.to_real a = ennreal.to_real b ↔ a = b :=
11181132
begin
11191133
rw ennreal.lt_top_iff_ne_top at *,
@@ -1124,7 +1138,7 @@ begin
11241138
{ assume h, rw h }
11251139
end
11261140

1127-
lemma to_real_mul_to_real {a b : ennreal} :
1141+
lemma to_real_mul_to_real :
11281142
(ennreal.to_real a) * (ennreal.to_real b) = ennreal.to_real (a * b) :=
11291143
begin
11301144
by_cases ha : a = ⊤,

src/data/set/basic.lean

Lines changed: 17 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -143,6 +143,9 @@ end set_coe
143143
/-- See also `subtype.prop` -/
144144
lemma subtype.mem {α : Type*} {s : set α} (p : s) : (p : α) ∈ s := p.prop
145145

146+
lemma eq.subset {α} {s t : set α} : s = t → s ⊆ t :=
147+
by { rintro rfl x hx, exact hx }
148+
146149
namespace set
147150

148151
variables {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {a : α} {s t : set α}
@@ -922,6 +925,10 @@ ext $ assume x, ⟨assume ⟨hx, _⟩, hx, assume h, ⟨h, not_false⟩⟩
922925
theorem diff_diff {u : set α} : s \ t \ u = s \ (t ∪ u) :=
923926
ext $ by simp [not_or_distrib, and.comm, and.left_comm]
924927

928+
-- the following statement contains parentheses to help the reader
929+
lemma diff_diff_comm {s t u : set α} : (s \ t) \ u = (s \ u) \ t :=
930+
by simp_rw [diff_diff, union_comm]
931+
925932
lemma diff_subset_iff {s t u : set α} : s \ t ⊆ u ↔ s ⊆ t ∪ u :=
926933
⟨assume h x xs, classical.by_cases or.inl (assume nxt, or.inr (h ⟨xs, nxt⟩)),
927934
assume h x ⟨xs, nxt⟩, or.resolve_left (h xs) nxt⟩
@@ -1327,7 +1334,7 @@ lemma image_preimage_inter (f : α → β) (s : set α) (t : set β) :
13271334
f '' (f ⁻¹' t ∩ s) = t ∩ f '' s :=
13281335
by simp only [inter_comm, image_inter_preimage]
13291336

1330-
lemma image_compl_preimage {f : α → β} {s : set α} {t : set β} : f '' (s \ f ⁻¹' t) = f '' s \ t :=
1337+
lemma image_diff_preimage {f : α → β} {s : set α} {t : set β} : f '' (s \ f ⁻¹' t) = f '' s \ t :=
13311338
by simp_rw [diff_eq, ← preimage_compl, image_inter_preimage]
13321339

13331340
theorem compl_image : image (compl : set α → set α) = preimage compl :=
@@ -1490,6 +1497,9 @@ lemma image_preimage_eq_of_subset {f : α → β} {s : set β} (hs : s ⊆ range
14901497
f '' (f ⁻¹' s) = s :=
14911498
by rw [image_preimage_eq_inter_range, inter_eq_self_of_subset_left hs]
14921499

1500+
lemma image_preimage_eq_iff {f : α → β} {s : set β} : f '' (f ⁻¹' s) = s ↔ s ⊆ range f :=
1501+
by { intro h, rw [← h], apply image_subset_range }, image_preimage_eq_of_subset⟩
1502+
14931503
lemma preimage_subset_preimage_iff {s t : set α} {f : β → α} (hs : s ⊆ range f) :
14941504
f ⁻¹' s ⊆ f ⁻¹' t ↔ s ⊆ t :=
14951505
begin
@@ -1538,6 +1548,12 @@ theorem preimage_singleton_eq_empty {f : α → β} {y : β} :
15381548
f ⁻¹' {y} = ∅ ↔ y ∉ range f :=
15391549
not_nonempty_iff_eq_empty.symm.trans $ not_congr preimage_singleton_nonempty
15401550

1551+
lemma range_subset_singleton {f : ι → α} {x : α} : range f ⊆ {x} ↔ f = const ι x :=
1552+
by simp [range_subset_iff, funext_iff, mem_singleton]
1553+
1554+
lemma image_compl_preimage {f : α → β} {s : set β} : f '' ((f ⁻¹' s)ᶜ) = range f \ s :=
1555+
by rw [compl_eq_univ_diff, image_diff_preimage, image_univ]
1556+
15411557
@[simp] theorem range_sigma_mk {β : α → Type*} (a : α) :
15421558
range (sigma.mk a : β a → Σ a, β a) = sigma.fst ⁻¹' {a} :=
15431559
begin

src/data/set/function.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -480,6 +480,8 @@ by simp [piecewise]
480480

481481
variable [∀j, decidable (j ∈ s)]
482482

483+
instance compl.decidable_mem (j : α) : decidable (j ∈ sᶜ) := not.decidable
484+
483485
lemma piecewise_insert [decidable_eq α] (j : α) [∀i, decidable (i ∈ insert j s)] :
484486
(insert j s).piecewise f g = function.update (s.piecewise f g) j (f j) :=
485487
begin
@@ -521,6 +523,20 @@ lemma piecewise_preimage (f g : α → β) (t) :
521523
s.piecewise f g ⁻¹' t = s ∩ f ⁻¹' t ∪ sᶜ ∩ g ⁻¹' t :=
522524
ext $ λ x, by by_cases x ∈ s; simp *
523525

526+
lemma comp_piecewise (h : β → γ) {f g : α → β} {x : α} :
527+
h (s.piecewise f g x) = s.piecewise (h ∘ f) (h ∘ g) x :=
528+
by by_cases hx : x ∈ s; simp [hx]
529+
530+
@[simp] lemma piecewise_same : s.piecewise f f = f :=
531+
by { ext x, by_cases hx : x ∈ s; simp [hx] }
532+
533+
lemma range_piecewise (f g : α → β) : range (s.piecewise f g) = f '' s ∪ g '' sᶜ :=
534+
begin
535+
ext y, split,
536+
{ rintro ⟨x, rfl⟩, by_cases h : x ∈ s;[left, right]; use x; simp [h] },
537+
{ rintro (⟨x, hx, rfl⟩|⟨x, hx, rfl⟩); use x; simp * at * }
538+
end
539+
524540
end set
525541

526542
namespace function

src/logic/function/basic.lean

Lines changed: 15 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -15,14 +15,23 @@ universes u v w
1515
namespace function
1616

1717
section
18-
variables: Sort u} {β : Sort v} {f : α → β}
18+
variablesβ γ : Sort*} {f : α → β}
1919

2020
/-- Evaluate a function at an argument. Useful if you want to talk about the partially applied
2121
`function.eval x : (Π x, β x) → β x`. -/
2222
@[reducible] def eval {β : α → Sort*} (x : α) (f : Π x, β x) : β x := f x
2323

2424
@[simp] lemma eval_apply {β : α → Sort*} (x : α) (f : Π x, β x) : eval x f = f x := rfl
2525

26+
lemma comp_apply {α : Sort u} {β : Sort v} {φ : Sort w} (f : β → φ) (g : α → β) (a : α) :
27+
(f ∘ g) a = f (g a) := rfl
28+
29+
@[simp] lemma const_apply {y : β} {x : α} : const α y x = y := rfl
30+
31+
@[simp] lemma const_comp {f : α → β} {c : γ} : const β c ∘ f = const α c := rfl
32+
33+
@[simp] lemma comp_const {f : β → γ} {b : β} : f ∘ const α b = const α (f b) := rfl
34+
2635
lemma hfunext {α α': Sort u} {β : α → Sort v} {β' : α' → Sort v} {f : Πa, β a} {f' : Πa, β' a}
2736
(hα : α = α') (h : ∀a a', a == a' → f a == f' a') : f == f' :=
2837
begin
@@ -40,9 +49,6 @@ end
4049
lemma funext_iff {β : α → Sort*} {f₁ f₂ : Π (x : α), β x} : f₁ = f₂ ↔ (∀a, f₁ a = f₂ a) :=
4150
iff.intro (assume h a, h ▸ rfl) funext
4251

43-
lemma comp_apply {α : Sort u} {β : Sort v} {φ : Sort w} (f : β → φ) (g : α → β) (a : α) :
44-
(f ∘ g) a = f (g a) := rfl
45-
4652
@[simp] theorem injective.eq_iff (I : injective f) {a b : α} :
4753
f a = f b ↔ a = b :=
4854
⟨@I _ _, congr_arg f⟩
@@ -66,10 +72,10 @@ the domain `α` also has decidable equality. -/
6672
def injective.decidable_eq [decidable_eq β] (I : injective f) : decidable_eq α :=
6773
λ a b, decidable_of_iff _ I.eq_iff
6874

69-
lemma injective.of_comp {γ : Sort w} {g : γ → α} (I : injective (f ∘ g)) : injective g :=
75+
lemma injective.of_comp {g : γ → α} (I : injective (f ∘ g)) : injective g :=
7076
λ x y h, I $ show f (g x) = f (g y), from congr_arg f h
7177

72-
lemma surjective.of_comp {γ : Sort w} {g : γ → α} (S : surjective (f ∘ g)) : surjective f :=
78+
lemma surjective.of_comp {g : γ → α} (S : surjective (f ∘ g)) : surjective f :=
7379
λ y, let ⟨x, h⟩ := S y in ⟨g x, h⟩
7480

7581
instance decidable_eq_pfun (p : Prop) [decidable p] (α : p → Type*)
@@ -140,11 +146,11 @@ funext h
140146
theorem right_inverse_iff_comp {f : α → β} {g : β → α} : right_inverse f g ↔ g ∘ f = id :=
141147
⟨right_inverse.comp_eq_id, congr_fun⟩
142148

143-
theorem left_inverse.comp {γ} {f : α → β} {g : β → α} {h : β → γ} {i : γ → β}
149+
theorem left_inverse.comp {f : α → β} {g : β → α} {h : β → γ} {i : γ → β}
144150
(hf : left_inverse f g) (hh : left_inverse h i) : left_inverse (h ∘ f) (g ∘ i) :=
145151
assume a, show h (f (g (i a))) = a, by rw [hf (i a), hh a]
146152

147-
theorem right_inverse.comp {γ} {f : α → β} {g : β → α} {h : β → γ} {i : γ → β}
153+
theorem right_inverse.comp {f : α → β} {g : β → α} {h : β → γ} {i : γ → β}
148154
(hf : right_inverse f g) (hh : right_inverse h i) : right_inverse (h ∘ f) (g ∘ i) :=
149155
left_inverse.comp hh hf
150156

@@ -346,7 +352,7 @@ lemma uncurry_def {α β γ} (f : α → β → γ) : uncurry f = (λp, f p.1 p.
346352
rfl
347353

348354
section bicomp
349-
variables: Type*} {β : Type*} {γ : Type*} {δ : Type*} {ε : Type*}
355+
variablesβ γ δ ε : Type*}
350356

351357
/-- Compose a binary function `f` with a pair of unary functions `g` and `h`.
352358
If both arguments of `f` have the same type and `g = h`, then `bicompl f g g = f on g`. -/

src/measure_theory/bochner_integration.lean

Lines changed: 5 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -69,9 +69,13 @@ The Bochner integral is defined following these steps:
6969
Some tips on how to prove a proposition if the API for the Bochner integral is not enough so that
7070
you need to unfold the definition of the Bochner integral and go back to simple functions.
7171
72+
One method is to use the theorem `integrable.induction` in the file `set_integral`, which allows
73+
you to prove something for an arbitrary measurable + integrable function.
74+
75+
Another method is using the following steps.
7276
See `integral_eq_lintegral_max_sub_lintegral_min` for a complicated example, which proves that
7377
`∫ f = ∫⁻ f⁺ - ∫⁻ f⁻`, with the first integral sign being the Bochner integral of a real-valued
74-
function f : α → ℝ, and second and third integral sign being the integral on ennreal-valued
78+
function `f : α → ℝ`, and second and third integral sign being the integral on ennreal-valued
7579
functions (called `lintegral`). The proof of `integral_eq_lintegral_max_sub_lintegral_min` is
7680
scattered in sections with the name `pos_part`.
7781

0 commit comments

Comments
 (0)