Skip to content

Commit 0f5133e

Browse files
committed
feat: add Set.IccExtend_eq_self (#4454)
Partial forward-port of leanprover-community/mathlib3#19097 Also rename `Set.Icc_extend_coe` to `Set.IccExtend_val`.
1 parent 115b6c3 commit 0f5133e

File tree

4 files changed

+19
-6
lines changed

4 files changed

+19
-6
lines changed

Mathlib/Analysis/SpecialFunctions/Trigonometric/Inverse.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -57,7 +57,7 @@ theorem neg_pi_div_two_le_arcsin (x : ℝ) : -(π / 2) ≤ arcsin x :=
5757
#align real.neg_pi_div_two_le_arcsin Real.neg_pi_div_two_le_arcsin
5858

5959
theorem arcsin_projIcc (x : ℝ) : arcsin (projIcc (-1) 1 (neg_le_self zero_le_one) x) = arcsin x :=
60-
by rw [arcsin, Function.comp_apply, Icc_extend_coe, Function.comp_apply, IccExtend,
60+
by rw [arcsin, Function.comp_apply, IccExtend_val, Function.comp_apply, IccExtend,
6161
Function.comp_apply]
6262
#align real.arcsin_proj_Icc Real.arcsin_projIcc
6363

Mathlib/Data/Set/Intervals/ProjIcc.lean

Lines changed: 16 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Yury G. Kudryashov, Patrick Massot
55
66
! This file was ported from Lean 3 source module data.set.intervals.proj_Icc
7-
! leanprover-community/mathlib commit aba57d4d3dae35460225919dcd82fe91355162f9
7+
! leanprover-community/mathlib commit 42e9a1fd3a99e10f82830349ba7f4f10e8961c2a
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -135,9 +135,22 @@ theorem IccExtend_of_mem (f : Icc a b → β) (hx : x ∈ Icc a b) : IccExtend h
135135
#align set.Icc_extend_of_mem Set.IccExtend_of_mem
136136

137137
@[simp]
138-
theorem Icc_extend_coe (f : Icc a b → β) (x : Icc a b) : IccExtend h f x = f x :=
138+
theorem IccExtend_val (f : Icc a b → β) (x : Icc a b) : IccExtend h f x = f x :=
139139
congr_arg f <| projIcc_val h x
140-
#align set.Icc_extend_coe Set.Icc_extend_coe
140+
#align set.Icc_extend_coe Set.IccExtend_val
141+
142+
/-- If `f : α → β` is a constant both on $(-∞, a]$ and on $[b, +∞)$, then the extension of this
143+
function from $[a, b]$ to the whole line is equal to the original function. -/
144+
theorem IccExtend_eq_self (f : α → β) (ha : ∀ x < a, f x = f a) (hb : ∀ x, b < x → f x = f b) :
145+
IccExtend h (f ∘ (↑)) = f := by
146+
ext x
147+
cases' lt_or_le x a with hxa hax
148+
· simp [IccExtend_of_le_left _ _ hxa.le, ha x hxa]
149+
· cases' le_or_lt x b with hxb hbx
150+
· lift x to Icc a b using ⟨hax, hxb⟩
151+
rw [IccExtend_val, comp_apply]
152+
· simp [IccExtend_of_right_le _ _ hbx.le, hb x hbx]
153+
#align set.Icc_extend_eq_self Set.IccExtend_eq_self
141154

142155
end Set
143156

Mathlib/Topology/Homotopy/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -185,7 +185,7 @@ theorem extend_apply_of_one_le (F : Homotopy f₀ f₁) {t : ℝ} (ht : 1 ≤ t)
185185

186186
@[simp]
187187
theorem extend_apply_coe (F : Homotopy f₀ f₁) (t : I) (x : X) : F.extend t x = F (t, x) :=
188-
ContinuousMap.congr_fun (Set.Icc_extend_coe (zero_le_one' ℝ) F.curry t) x
188+
ContinuousMap.congr_fun (Set.IccExtend_val (zero_le_one' ℝ) F.curry t) x
189189
#align continuous_map.homotopy.extend_apply_coe ContinuousMap.Homotopy.extend_apply_coe
190190

191191
@[simp]

Mathlib/Topology/PathConnected.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -277,7 +277,7 @@ theorem extend_one : γ.extend 1 = y := by simp
277277
@[simp]
278278
theorem extend_extends' {X : Type _} [TopologicalSpace X] {a b : X} (γ : Path a b)
279279
(t : (Icc 0 1 : Set ℝ)) : γ.extend t = γ t :=
280-
Icc_extend_coe _ γ t
280+
IccExtend_val _ γ t
281281
#align path.extend_extends' Path.extend_extends'
282282

283283
@[simp]

0 commit comments

Comments
 (0)