@@ -22,14 +22,11 @@ lemma surj_on_Ioo_of_monotone_surjective
22
22
(h_mono : monotone f) (h_surj : function.surjective f) (a b : α) :
23
23
surj_on f (Ioo a b) (Ioo (f a) (f b)) :=
24
24
begin
25
- classical,
26
25
intros p hp,
27
26
rcases h_surj p with ⟨x, rfl⟩,
28
27
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 )
33
30
end
34
31
35
32
lemma surj_on_Ico_of_monotone_surjective
@@ -38,12 +35,10 @@ lemma surj_on_Ico_of_monotone_surjective
38
35
begin
39
36
obtain hab | hab := lt_or_le a b,
40
37
{ 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) },
44
40
{ 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 } },
47
42
{ rw Ico_eq_empty (h_mono hab).not_lt,
48
43
exact surj_on_empty f _ }
49
44
end
@@ -58,32 +53,21 @@ lemma surj_on_Icc_of_monotone_surjective
58
53
(h_mono : monotone f) (h_surj : function.surjective f) {a b : α} (hab : a ≤ b) :
59
54
surj_on f (Icc a b) (Icc (f a) (f b)) :=
60
55
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 }
74
62
end
75
63
76
64
lemma surj_on_Ioi_of_monotone_surjective
77
65
(h_mono : monotone f) (h_surj : function.surjective f) (a : α) :
78
66
surj_on f (Ioi a) (Ioi (f a)) :=
79
67
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
87
71
end
88
72
89
73
lemma surj_on_Iio_of_monotone_surjective
@@ -95,13 +79,9 @@ lemma surj_on_Ici_of_monotone_surjective
95
79
(h_mono : monotone f) (h_surj : function.surjective f) (a : α) :
96
80
surj_on f (Ici a) (Ici (f a)) :=
97
81
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 _ _)
105
85
end
106
86
107
87
lemma surj_on_Iic_of_monotone_surjective
0 commit comments