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

Commit a370cf0

Browse files
committed
chore(data/set/intervals): golf, rename (#12893)
* rename `set.mem_Ioo_or_eq_endpoints_of_mem_Icc` → `set.eq_endpoints_or_mem_Ioo_of_mem_Icc`; * rename `set.mem_Ioo_or_eq_left_of_mem_Ico` → `set.eq_left_or_mem_Ioo_of_mem_Ico`; * rename `set.mem_Ioo_or_eq_right_of_mem_Ioc` → `set.eq_right_or_mem_Ioo_of_mem_Ioc`; * golf the proofs of these lemmas. The new names better reflect the order of terms in `or`.
1 parent 5ef365a commit a370cf0

File tree

3 files changed

+27
-62
lines changed

3 files changed

+27
-62
lines changed

src/data/set/intervals/basic.lean

Lines changed: 8 additions & 23 deletions
Original file line numberDiff line numberDiff line change
@@ -484,32 +484,17 @@ begin
484484
apply_rules [subset_diff_singleton] }
485485
end
486486

487-
lemma mem_Ioo_or_eq_endpoints_of_mem_Icc {x : α} (hmem : x ∈ Icc a b) :
488-
x = a ∨ x = b ∨ x ∈ Ioo a b :=
489-
begin
490-
rw [mem_Icc, le_iff_lt_or_eq, le_iff_lt_or_eq] at hmem,
491-
rcases hmem with ⟨hxa | hxa, hxb | hxb⟩,
492-
{ exact or.inr (or.inr ⟨hxa, hxb⟩) },
493-
{ exact or.inr (or.inl hxb) },
494-
all_goals { exact or.inl hxa.symm }
495-
end
496-
497-
lemma mem_Ioo_or_eq_left_of_mem_Ico {x : α} (hmem : x ∈ Ico a b) :
487+
lemma eq_left_or_mem_Ioo_of_mem_Ico {x : α} (hmem : x ∈ Ico a b) :
498488
x = a ∨ x ∈ Ioo a b :=
499-
begin
500-
rw [mem_Ico, le_iff_lt_or_eq] at hmem,
501-
rcases hmem with ⟨hxa | hxa, hxb⟩,
502-
{ exact or.inr ⟨hxa, hxb⟩ },
503-
{ exact or.inl hxa.symm }
504-
end
489+
hmem.1.eq_or_gt.imp_right $ λ h, ⟨h, hmem.2
505490

506-
lemma mem_Ioo_or_eq_right_of_mem_Ioc {x : α} (hmem : x ∈ Ioc a b) :
491+
lemma eq_right_or_mem_Ioo_of_mem_Ioc {x : α} (hmem : x ∈ Ioc a b) :
507492
x = b ∨ x ∈ Ioo a b :=
508-
begin
509-
have := @mem_Ioo_or_eq_left_of_mem_Ico _ _ (to_dual b) (to_dual a) (to_dual x),
510-
rw [dual_Ioo, dual_Ico] at this,
511-
exact this hmem
512-
end
493+
hmem.2.eq_or_lt.imp_right $ and.intro hmem.1
494+
495+
lemma eq_endpoints_or_mem_Ioo_of_mem_Icc {x : α} (hmem : x ∈ Icc a b) :
496+
x = a ∨ x = b ∨ x ∈ Ioo a b :=
497+
hmem.1.eq_or_gt.imp_right $ λ h, eq_right_or_mem_Ioo_of_mem_Ioc ⟨h, hmem.2
513498

514499
lemma _root_.is_max.Ici_eq (h : is_max a) : Ici a = {a} :=
515500
eq_singleton_iff_unique_mem.2 ⟨left_mem_Ici, λ b, h.eq_of_ge⟩

src/data/set/intervals/surj_on.lean

Lines changed: 17 additions & 37 deletions
Original file line numberDiff line numberDiff line change
@@ -22,14 +22,11 @@ lemma surj_on_Ioo_of_monotone_surjective
2222
(h_mono : monotone f) (h_surj : function.surjective f) (a b : α) :
2323
surj_on f (Ioo a b) (Ioo (f a) (f b)) :=
2424
begin
25-
classical,
2625
intros p hp,
2726
rcases h_surj p with ⟨x, rfl⟩,
2827
refine ⟨x, mem_Ioo.2 _, rfl⟩,
29-
by_contra h,
30-
cases not_and_distrib.mp h with ha hb,
31-
{ exact has_lt.lt.false (lt_of_lt_of_le hp.1 (h_mono (not_lt.mp ha))) },
32-
{ exact has_lt.lt.false (lt_of_le_of_lt (h_mono (not_lt.mp hb)) hp.2) }
28+
contrapose! hp,
29+
exact λ h, h.2.not_le (h_mono $ hp $ h_mono.reflect_lt h.1)
3330
end
3431

3532
lemma surj_on_Ico_of_monotone_surjective
@@ -38,12 +35,10 @@ lemma surj_on_Ico_of_monotone_surjective
3835
begin
3936
obtain hab | hab := lt_or_le a b,
4037
{ intros p hp,
41-
rcases mem_Ioo_or_eq_left_of_mem_Ico hp with hp'|hp',
42-
{ rw hp',
43-
exact ⟨a, left_mem_Ico.mpr hab, rfl⟩ },
38+
rcases eq_left_or_mem_Ioo_of_mem_Ico hp with rfl|hp',
39+
{ exact mem_image_of_mem f (left_mem_Ico.mpr hab) },
4440
{ have := surj_on_Ioo_of_monotone_surjective h_mono h_surj a b hp',
45-
cases this with x hx,
46-
exact ⟨x, Ioo_subset_Ico_self hx.1, hx.2⟩ } },
41+
exact image_subset f Ioo_subset_Ico_self this } },
4742
{ rw Ico_eq_empty (h_mono hab).not_lt,
4843
exact surj_on_empty f _ }
4944
end
@@ -58,32 +53,21 @@ lemma surj_on_Icc_of_monotone_surjective
5853
(h_mono : monotone f) (h_surj : function.surjective f) {a b : α} (hab : a ≤ b) :
5954
surj_on f (Icc a b) (Icc (f a) (f b)) :=
6055
begin
61-
rcases lt_or_eq_of_le hab with hab|hab,
62-
{ intros p hp,
63-
rcases mem_Ioo_or_eq_endpoints_of_mem_Icc hp with hp'|⟨hp'|hp'⟩,
64-
{ rw hp',
65-
refine ⟨a, left_mem_Icc.mpr (le_of_lt hab), rfl⟩ },
66-
{ rw hp',
67-
refine ⟨b, right_mem_Icc.mpr (le_of_lt hab), rfl⟩ },
68-
{ have := surj_on_Ioo_of_monotone_surjective h_mono h_surj a b hp',
69-
cases this with x hx,
70-
exact ⟨x, Ioo_subset_Icc_self hx.1, hx.2⟩ } },
71-
{ simp only [hab, Icc_self],
72-
intros _ hp,
73-
exact ⟨b, mem_singleton _, (mem_singleton_iff.mp hp).symm⟩ }
56+
intros p hp,
57+
rcases eq_endpoints_or_mem_Ioo_of_mem_Icc hp with (rfl|rfl|hp'),
58+
{ exact ⟨a, left_mem_Icc.mpr hab, rfl⟩ },
59+
{ exact ⟨b, right_mem_Icc.mpr hab, rfl⟩ },
60+
{ have := surj_on_Ioo_of_monotone_surjective h_mono h_surj a b hp',
61+
exact image_subset f Ioo_subset_Icc_self this }
7462
end
7563

7664
lemma surj_on_Ioi_of_monotone_surjective
7765
(h_mono : monotone f) (h_surj : function.surjective f) (a : α) :
7866
surj_on f (Ioi a) (Ioi (f a)) :=
7967
begin
80-
classical,
81-
intros p hp,
82-
rcases h_surj p with ⟨x, rfl⟩,
83-
refine ⟨x, _, rfl⟩,
84-
simp only [mem_Ioi],
85-
by_contra h,
86-
exact has_lt.lt.false (lt_of_lt_of_le hp (h_mono (not_lt.mp h)))
68+
rw [← compl_Iic, ← compl_compl (Ioi (f a))],
69+
refine maps_to.surj_on_compl _ h_surj,
70+
exact λ x hx, (h_mono hx).not_lt
8771
end
8872

8973
lemma surj_on_Iio_of_monotone_surjective
@@ -95,13 +79,9 @@ lemma surj_on_Ici_of_monotone_surjective
9579
(h_mono : monotone f) (h_surj : function.surjective f) (a : α) :
9680
surj_on f (Ici a) (Ici (f a)) :=
9781
begin
98-
intros p hp,
99-
rw [mem_Ici, le_iff_lt_or_eq] at hp,
100-
rcases hp with hp'|hp',
101-
{ cases (surj_on_Ioi_of_monotone_surjective h_mono h_surj a hp') with x hx,
102-
exact ⟨x, Ioi_subset_Ici_self hx.1, hx.2⟩ },
103-
{ rw ← hp',
104-
refine ⟨a, left_mem_Ici, rfl⟩ }
82+
rw [← Ioi_union_left, ← Ioi_union_left],
83+
exact (surj_on_Ioi_of_monotone_surjective h_mono h_surj a).union_union
84+
(@image_singleton _ _ f a ▸ surj_on_image _ _)
10585
end
10686

10787
lemma surj_on_Iic_of_monotone_surjective

src/topology/algebra/order/extend_from.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -25,7 +25,7 @@ begin
2525
apply continuous_on_extend_from,
2626
{ rw closure_Ioo hab },
2727
{ intros x x_in,
28-
rcases mem_Ioo_or_eq_endpoints_of_mem_Icc x_in with rfl | rfl | h,
28+
rcases eq_endpoints_or_mem_Ioo_of_mem_Icc x_in with rfl | rfl | h,
2929
{ exact ⟨la, ha.mono_left $ nhds_within_mono _ Ioo_subset_Ioi_self⟩ },
3030
{ exact ⟨lb, hb.mono_left $ nhds_within_mono _ Ioo_subset_Iio_self⟩ },
3131
{ use [f x, hf x h] } }
@@ -62,7 +62,7 @@ begin
6262
apply continuous_on_extend_from,
6363
{ rw [closure_Ioo hab.ne], exact Ico_subset_Icc_self, },
6464
{ intros x x_in,
65-
rcases mem_Ioo_or_eq_left_of_mem_Ico x_in with rfl | h,
65+
rcases eq_left_or_mem_Ioo_of_mem_Ico x_in with rfl | h,
6666
{ use la,
6767
simpa [hab] },
6868
{ use [f x, hf x h] } }

0 commit comments

Comments
 (0)