@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Yury G. Kudryashov, Patrick Massot
5
5
6
6
! 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
8
8
! Please do not edit these lines, except to modify the commit id
9
9
! if you have ported upstream changes.
10
10
-/
@@ -135,9 +135,22 @@ theorem IccExtend_of_mem (f : Icc a b → β) (hx : x ∈ Icc a b) : IccExtend h
135
135
#align set.Icc_extend_of_mem Set.IccExtend_of_mem
136
136
137
137
@[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 :=
139
139
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
141
154
142
155
end Set
143
156
0 commit comments