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

Commit 5cc2dfd

Browse files
committed
feat(*): add various results split out from proof of Gallagher's theorem (#17985)
This is a collection of tiny results that are not really related but I'm bundling them together in an effort to aid review.
1 parent 3d95492 commit 5cc2dfd

File tree

9 files changed

+174
-41
lines changed

9 files changed

+174
-41
lines changed

src/analysis/normed/field/basic.lean

Lines changed: 10 additions & 27 deletions
Original file line numberDiff line numberDiff line change
@@ -693,6 +693,8 @@ instance : normed_comm_ring ℤ :=
693693

694694
lemma int.norm_eq_abs (n : ℤ) : ‖n‖ = |n| := rfl
695695

696+
@[simp] lemma int.norm_coe_nat (n : ℕ) : ‖(n : ℤ)‖ = n := by simp [int.norm_eq_abs]
697+
696698
lemma nnreal.coe_nat_abs (n : ℤ) : (n.nat_abs : ℝ≥0) = ‖n‖₊ :=
697699
nnreal.eq $ calc ((n.nat_abs : ℝ≥0) : ℝ)
698700
= (n.nat_abs : ℤ) : by simp only [int.cast_coe_nat, nnreal.coe_nat_cast]
@@ -724,36 +726,17 @@ instance : densely_normed_field ℚ :=
724726
by rw [← rat.norm_cast_real, ← int.norm_cast_real]; congr' 1; norm_cast
725727

726728
-- Now that we've installed the norm on `ℤ`,
727-
-- we can state some lemmas about `nsmul` and `zsmul`.
729+
-- we can state some lemmas about `zsmul`.
728730
section
729-
variables [seminormed_add_comm_group α]
730-
731-
lemma norm_nsmul_le (n : ℕ) (a : α) : ‖n • a‖ ≤ n * ‖a‖ :=
732-
begin
733-
induction n with n ih,
734-
{ simp only [norm_zero, nat.cast_zero, zero_mul, zero_smul] },
735-
simp only [nat.succ_eq_add_one, add_smul, add_mul, one_mul, nat.cast_add,
736-
nat.cast_one, one_nsmul],
737-
exact norm_add_le_of_le ih le_rfl
738-
end
739-
740-
lemma norm_zsmul_le (n : ℤ) (a : α) : ‖n • a‖ ≤ ‖n‖ * ‖a‖ :=
741-
begin
742-
induction n with n n,
743-
{ simp only [int.of_nat_eq_coe, coe_nat_zsmul],
744-
convert norm_nsmul_le n a,
745-
exact nat.abs_cast n },
746-
{ simp only [int.neg_succ_of_nat_coe, neg_smul, norm_neg, coe_nat_zsmul],
747-
convert norm_nsmul_le n.succ a,
748-
exact nat.abs_cast n.succ, }
749-
end
731+
variables [seminormed_comm_group α]
750732

751-
lemma nnnorm_nsmul_le (n : ℕ) (a : α) : ‖n • a‖₊ ≤ n * ‖a‖₊ :=
752-
by simpa only [←nnreal.coe_le_coe, nnreal.coe_mul, nnreal.coe_nat_cast]
753-
using norm_nsmul_le n a
733+
@[to_additive norm_zsmul_le]
734+
lemma norm_zpow_le_mul_norm (n : ℤ) (a : α) : ‖a^n‖ ≤ ‖n‖ * ‖a‖ :=
735+
by rcases n.eq_coe_or_neg with ⟨n, rfl | rfl⟩; simpa using norm_pow_le_mul_norm n a
754736

755-
lemma nnnorm_zsmul_le (n : ℤ) (a : α) : ‖n • a‖₊ ≤ ‖n‖₊ * ‖a‖₊ :=
756-
by simpa only [←nnreal.coe_le_coe, nnreal.coe_mul] using norm_zsmul_le n a
737+
@[to_additive nnnorm_zsmul_le]
738+
lemma nnnorm_zpow_le_mul_norm (n : ℤ) (a : α) : ‖a^n‖₊ ≤ ‖n‖₊ * ‖a‖₊ :=
739+
by simpa only [← nnreal.coe_le_coe, nnreal.coe_mul] using norm_zpow_le_mul_norm n a
757740

758741
end
759742

src/analysis/normed/group/basic.lean

Lines changed: 36 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1061,6 +1061,42 @@ by { ext c,
10611061
((*) b) ⁻¹' sphere a r = sphere (a / b) r :=
10621062
by { ext c, simp only [set.mem_preimage, mem_sphere_iff_norm', div_div_eq_mul_div, mul_comm] }
10631063

1064+
@[to_additive norm_nsmul_le] lemma norm_pow_le_mul_norm (n : ℕ) (a : E) : ‖a^n‖ ≤ n * ‖a‖ :=
1065+
begin
1066+
induction n with n ih, { simp, },
1067+
simpa only [pow_succ', nat.cast_succ, add_mul, one_mul] using norm_mul_le_of_le ih le_rfl,
1068+
end
1069+
1070+
@[to_additive nnnorm_nsmul_le] lemma nnnorm_pow_le_mul_norm (n : ℕ) (a : E) : ‖a^n‖₊ ≤ n * ‖a‖₊ :=
1071+
by simpa only [← nnreal.coe_le_coe, nnreal.coe_mul, nnreal.coe_nat_cast]
1072+
using norm_pow_le_mul_norm n a
1073+
1074+
@[to_additive] lemma pow_mem_closed_ball {n : ℕ} (h : a ∈ closed_ball b r) :
1075+
a^n ∈ closed_ball (b^n) (n • r) :=
1076+
begin
1077+
simp only [mem_closed_ball, dist_eq_norm_div, ← div_pow] at h ⊢,
1078+
refine (norm_pow_le_mul_norm n (a / b)).trans _,
1079+
simpa only [nsmul_eq_mul] using mul_le_mul_of_nonneg_left h n.cast_nonneg,
1080+
end
1081+
1082+
@[to_additive] lemma pow_mem_ball {n : ℕ} (hn : 0 < n) (h : a ∈ ball b r) :
1083+
a^n ∈ ball (b^n) (n • r) :=
1084+
begin
1085+
simp only [mem_ball, dist_eq_norm_div, ← div_pow] at h ⊢,
1086+
refine lt_of_le_of_lt (norm_pow_le_mul_norm n (a / b)) _,
1087+
replace hn : 0 < (n : ℝ), { norm_cast, assumption, },
1088+
rw nsmul_eq_mul,
1089+
nlinarith,
1090+
end
1091+
1092+
@[simp, to_additive] lemma mul_mem_closed_ball_mul_iff {c : E} :
1093+
a * c ∈ closed_ball (b * c) r ↔ a ∈ closed_ball b r :=
1094+
by simp only [mem_closed_ball, dist_eq_norm_div, mul_div_mul_right_eq_div]
1095+
1096+
@[simp, to_additive] lemma mul_mem_ball_mul_iff {c : E} :
1097+
a * c ∈ ball (b * c) r ↔ a ∈ ball b r :=
1098+
by simp only [mem_ball, dist_eq_norm_div, mul_div_mul_right_eq_div]
1099+
10641100
namespace isometric
10651101

10661102
/-- Multiplication `y ↦ x * y` as an `isometry`. -/

src/analysis/normed_space/int.lean

Lines changed: 0 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -31,8 +31,6 @@ by rw [← coe_nnnorm, int.nnnorm_coe_units, nnreal.coe_one]
3131

3232
@[simp] lemma nnnorm_coe_nat (n : ℕ) : ‖(n : ℤ)‖₊ = n := real.nnnorm_coe_nat _
3333

34-
@[simp] lemma norm_coe_nat (n : ℕ) : ‖(n : ℤ)‖ = n := real.norm_coe_nat _
35-
3634
@[simp] lemma to_nat_add_to_nat_neg_eq_nnnorm (n : ℤ) : ↑(n.to_nat) + ↑((-n).to_nat) = ‖n‖₊ :=
3735
by rw [← nat.cast_add, to_nat_add_to_nat_neg_eq_nat_abs, nnreal.coe_nat_abs]
3836

src/data/nat/totient.lean

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -36,6 +36,17 @@ by simp [totient]
3636

3737
lemma totient_eq_card_coprime (n : ℕ) : φ n = ((range n).filter n.coprime).card := rfl
3838

39+
/-- A characterisation of `nat.totient` that avoids `finset`. -/
40+
lemma totient_eq_card_lt_and_coprime (n : ℕ) : φ n = nat.card {m | m < n ∧ n.coprime m} :=
41+
begin
42+
let e : {m | m < n ∧ n.coprime m} ≃ finset.filter n.coprime (finset.range n) :=
43+
{ to_fun := λ m, ⟨m, by simpa only [finset.mem_filter, finset.mem_range] using m.property⟩,
44+
inv_fun := λ m, ⟨m, by simpa only [finset.mem_filter, finset.mem_range] using m.property⟩,
45+
left_inv := λ m, by simp only [subtype.coe_mk, subtype.coe_eta],
46+
right_inv := λ m, by simp only [subtype.coe_mk, subtype.coe_eta] },
47+
rw [totient_eq_card_coprime, card_congr e, card_eq_fintype_card, fintype.card_coe],
48+
end
49+
3950
lemma totient_le (n : ℕ) : φ n ≤ n :=
4051
((range n).card_filter_le _).trans_eq (card_range n)
4152

src/measure_theory/covering/liminf_limsup.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -163,7 +163,7 @@ begin
163163
cases le_or_lt 1 M with hM' hM',
164164
{ apply has_subset.subset.eventually_le,
165165
change _ ≤ _,
166-
refine mono_blimsup' (hMr.mono $ λ i hi, cthickening_mono _ (s i)),
166+
refine mono_blimsup' (hMr.mono $ λ i hi hp, cthickening_mono _ (s i)),
167167
exact (le_mul_of_one_le_left (hRp i) hM').trans hi, },
168168
{ simp only [← @cthickening_closure _ _ _ (s _)],
169169
have hs : ∀ i, is_closed (closure (s i)) := λ i, is_closed_closure,

src/measure_theory/measurable_space.lean

Lines changed: 22 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -1458,18 +1458,36 @@ instance : boolean_algebra (subtype (measurable_set : set α → Prop)) :=
14581458
.. measurable_set.subtype.bounded_order,
14591459
.. measurable_set.subtype.distrib_lattice }
14601460

1461+
@[measurability] lemma measurable_set_blimsup {s : ℕ → set α} {p : ℕ → Prop}
1462+
(h : ∀ n, p n → measurable_set (s n)) :
1463+
measurable_set $ filter.blimsup s filter.at_top p :=
1464+
begin
1465+
simp only [filter.blimsup_eq_infi_bsupr_of_nat, supr_eq_Union, infi_eq_Inter],
1466+
exact measurable_set.Inter
1467+
(λ n, measurable_set.Union (λ m, measurable_set.Union $ λ hm, h m hm.1)),
1468+
end
1469+
1470+
@[measurability] lemma measurable_set_bliminf {s : ℕ → set α} {p : ℕ → Prop}
1471+
(h : ∀ n, p n → measurable_set (s n)) :
1472+
measurable_set $ filter.bliminf s filter.at_top p :=
1473+
begin
1474+
simp only [filter.bliminf_eq_supr_binfi_of_nat, infi_eq_Inter, supr_eq_Union],
1475+
exact measurable_set.Union
1476+
(λ n, measurable_set.Inter (λ m, measurable_set.Inter $ λ hm, h m hm.1)),
1477+
end
1478+
14611479
@[measurability] lemma measurable_set_limsup {s : ℕ → set α} (hs : ∀ n, measurable_set $ s n) :
14621480
measurable_set $ filter.limsup s filter.at_top :=
14631481
begin
1464-
simp only [filter.limsup_eq_infi_supr_of_nat', supr_eq_Union, infi_eq_Inter],
1465-
exact measurable_set.Inter (λ n, measurable_set.Union $ λ m, hs $ m + n),
1482+
convert measurable_set_blimsup (λ n h, hs n : ∀ n, true → measurable_set (s n)),
1483+
simp,
14661484
end
14671485

14681486
@[measurability] lemma measurable_set_liminf {s : ℕ → set α} (hs : ∀ n, measurable_set $ s n) :
14691487
measurable_set $ filter.liminf s filter.at_top :=
14701488
begin
1471-
simp only [filter.liminf_eq_supr_infi_of_nat', supr_eq_Union, infi_eq_Inter],
1472-
exact measurable_set.Union (λ n, measurable_set.Inter $ λ m, hs $ m + n),
1489+
convert measurable_set_bliminf (λ n h, hs n : ∀ n, true → measurable_set (s n)),
1490+
simp,
14731491
end
14741492

14751493
end measurable_set

src/measure_theory/measure/measure_space_def.lean

Lines changed: 32 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -406,6 +406,38 @@ lemma ae_eq_set_union {s' t' : set α} (h : s =ᵐ[μ] t) (h' : s' =ᵐ[μ] t')
406406
(s ∪ s' : set α) =ᵐ[μ] (t ∪ t' : set α) :=
407407
h.union h'
408408

409+
lemma union_ae_eq_univ_of_ae_eq_univ_left (h : s =ᵐ[μ] univ) :
410+
(s ∪ t : set α) =ᵐ[μ] univ :=
411+
by { convert ae_eq_set_union h (ae_eq_refl t), rw univ_union, }
412+
413+
lemma union_ae_eq_univ_of_ae_eq_univ_right (h : t =ᵐ[μ] univ) :
414+
(s ∪ t : set α) =ᵐ[μ] univ :=
415+
by { convert ae_eq_set_union (ae_eq_refl s) h, rw union_univ, }
416+
417+
lemma union_ae_eq_right_of_ae_eq_empty (h : s =ᵐ[μ] (∅ : set α)) :
418+
(s ∪ t : set α) =ᵐ[μ] t :=
419+
by { convert ae_eq_set_union h (ae_eq_refl t), rw empty_union, }
420+
421+
lemma union_ae_eq_left_of_ae_eq_empty (h : t =ᵐ[μ] (∅ : set α)) :
422+
(s ∪ t : set α) =ᵐ[μ] s :=
423+
by { convert ae_eq_set_union (ae_eq_refl s) h, rw union_empty, }
424+
425+
lemma inter_ae_eq_right_of_ae_eq_univ (h : s =ᵐ[μ] univ) :
426+
(s ∩ t : set α) =ᵐ[μ] t :=
427+
by { convert ae_eq_set_inter h (ae_eq_refl t), rw univ_inter, }
428+
429+
lemma inter_ae_eq_left_of_ae_eq_univ (h : t =ᵐ[μ] univ) :
430+
(s ∩ t : set α) =ᵐ[μ] s :=
431+
by { convert ae_eq_set_inter (ae_eq_refl s) h, rw inter_univ, }
432+
433+
lemma inter_ae_eq_empty_of_ae_eq_empty_left (h : s =ᵐ[μ] (∅ : set α)) :
434+
(s ∩ t : set α) =ᵐ[μ] (∅ : set α) :=
435+
by { convert ae_eq_set_inter h (ae_eq_refl t), rw empty_inter, }
436+
437+
lemma inter_ae_eq_empty_of_ae_eq_empty_right (h : t =ᵐ[μ] (∅ : set α)) :
438+
(s ∩ t : set α) =ᵐ[μ] (∅ : set α) :=
439+
by { convert ae_eq_set_inter (ae_eq_refl s) h, rw inter_empty, }
440+
409441
@[to_additive]
410442
lemma _root_.set.mul_indicator_ae_eq_one {M : Type*} [has_one M] {f : α → M} {s : set α}
411443
(h : s.mul_indicator f =ᵐ[μ] 1) : μ (s ∩ function.mul_support f) = 0 :=

src/order/hom/complete_lattice.lean

Lines changed: 30 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -583,7 +583,9 @@ end complete_lattice_hom
583583

584584
namespace complete_lattice_hom
585585

586-
/-- `set.preimage` as a complete lattice homomorphism. -/
586+
/-- `set.preimage` as a complete lattice homomorphism.
587+
588+
See also `Sup_hom.set_image`. -/
587589
def set_preimage (f : α → β) : complete_lattice_hom (set β) (set α) :=
588590
{ to_fun := preimage f,
589591
map_Sup' := λ s, preimage_sUnion.trans $ by simp only [set.Sup_eq_sUnion, set.sUnion_image],
@@ -597,3 +599,30 @@ lemma set_preimage_comp (g : β → γ) (f : α → β) :
597599
set_preimage (g ∘ f) = (set_preimage f).comp (set_preimage g) := rfl
598600

599601
end complete_lattice_hom
602+
603+
lemma set.image_Sup {f : α → β} (s : set (set α)) :
604+
f '' Sup s = Sup (image f '' s) :=
605+
begin
606+
ext b,
607+
simp only [Sup_eq_sUnion, mem_image, mem_sUnion, exists_prop, sUnion_image, mem_Union],
608+
split,
609+
{ rintros ⟨a, ⟨t, ht₁, ht₂⟩, rfl⟩, exact ⟨t, ht₁, a, ht₂, rfl⟩, },
610+
{ rintros ⟨t, ht₁, a, ht₂, rfl⟩, exact ⟨a, ⟨t, ht₁, ht₂⟩, rfl⟩, },
611+
end
612+
613+
/-- Using `set.image`, a function between types yields a `Sup_hom` between their lattices of
614+
subsets.
615+
616+
See also `complete_lattice_hom.set_preimage`. -/
617+
@[simps] def Sup_hom.set_image (f : α → β) : Sup_hom (set α) (set β) :=
618+
{ to_fun := image f,
619+
map_Sup' := set.image_Sup }
620+
621+
/-- An equivalence of types yields an order isomorphism between their lattices of subsets. -/
622+
@[simps] def equiv.to_order_iso_set (e : α ≃ β) : set α ≃o set β :=
623+
{ to_fun := image e,
624+
inv_fun := image e.symm,
625+
left_inv := λ s, by simp only [← image_comp, equiv.symm_comp_self, id.def, image_id'],
626+
right_inv := λ s, by simp only [← image_comp, equiv.self_comp_symm, id.def, image_id'],
627+
map_rel_iff' :=
628+
λ s t, ⟨λ h, by simpa using @monotone_image _ _ e.symm _ _ h, λ h, monotone_image h⟩ }

src/order/liminf_limsup.lean

Lines changed: 32 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -661,19 +661,19 @@ lemma bliminf_antitone (h : ∀ x, p x → q x) :
661661
bliminf u f q ≤ bliminf u f p :=
662662
Sup_le_Sup $ λ a ha, ha.mono $ by tauto
663663

664-
lemma mono_blimsup' (h : ∀ᶠ x in f, u x ≤ v x) :
664+
lemma mono_blimsup' (h : ∀ᶠ x in f, p x → u x ≤ v x) :
665665
blimsup u f p ≤ blimsup v f p :=
666-
Inf_le_Inf $ λ a ha, (ha.and h).mono $ λ x hx hx', hx.2.trans (hx.1 hx')
666+
Inf_le_Inf $ λ a ha, (ha.and h).mono $ λ x hx hx', (hx.2 hx').trans (hx.1 hx')
667667

668-
lemma mono_blimsup (h : ∀ x, u x ≤ v x) :
668+
lemma mono_blimsup (h : ∀ x, p x → u x ≤ v x) :
669669
blimsup u f p ≤ blimsup v f p :=
670670
mono_blimsup' $ eventually_of_forall h
671671

672-
lemma mono_bliminf' (h : ∀ᶠ x in f, u x ≤ v x) :
672+
lemma mono_bliminf' (h : ∀ᶠ x in f, p x → u x ≤ v x) :
673673
bliminf u f p ≤ bliminf v f p :=
674-
Sup_le_Sup $ λ a ha, (ha.and h).mono $ λ x hx hx', (hx.1 hx').trans hx.2
674+
Sup_le_Sup $ λ a ha, (ha.and h).mono $ λ x hx hx', (hx.1 hx').trans (hx.2 hx')
675675

676-
lemma mono_bliminf (h : ∀ x, u x ≤ v x) :
676+
lemma mono_bliminf (h : ∀ x, p x → u x ≤ v x) :
677677
bliminf u f p ≤ bliminf v f p :=
678678
mono_bliminf' $ eventually_of_forall h
679679

@@ -703,6 +703,32 @@ sup_le (blimsup_mono $ by tauto) (blimsup_mono $ by tauto)
703703
bliminf u f (λ x, p x ∨ q x) ≤ bliminf u f p ⊓ bliminf u f q :=
704704
@blimsup_sup_le_or αᵒᵈ β _ f p q u
705705

706+
lemma order_iso.apply_blimsup [complete_lattice γ] (e : α ≃o γ) :
707+
e (blimsup u f p) = blimsup (e ∘ u) f p :=
708+
begin
709+
simp only [blimsup_eq, map_Inf, function.comp_app],
710+
congr,
711+
ext c,
712+
obtain ⟨a, rfl⟩ := e.surjective c,
713+
simp,
714+
end
715+
716+
lemma order_iso.apply_bliminf [complete_lattice γ] (e : α ≃o γ) :
717+
e (bliminf u f p) = bliminf (e ∘ u) f p :=
718+
@order_iso.apply_blimsup αᵒᵈ β γᵒᵈ _ f p u _ e.dual
719+
720+
lemma Sup_hom.apply_blimsup_le [complete_lattice γ] (g : Sup_hom α γ) :
721+
g (blimsup u f p) ≤ blimsup (g ∘ u) f p :=
722+
begin
723+
simp only [blimsup_eq_infi_bsupr],
724+
refine ((order_hom_class.mono g).map_infi₂_le _).trans _,
725+
simp only [_root_.map_supr],
726+
end
727+
728+
lemma Inf_hom.le_apply_bliminf [complete_lattice γ] (g : Inf_hom α γ) :
729+
bliminf (g ∘ u) f p ≤ g (bliminf u f p) :=
730+
@Sup_hom.apply_blimsup_le αᵒᵈ β γᵒᵈ _ f p u _ g.dual
731+
706732
end complete_lattice
707733

708734
section complete_distrib_lattice

0 commit comments

Comments
 (0)