Skip to content

Commit 5fcd552

Browse files
committed
chore: gcongr tagging (#7889)
Tag lemmas about `tsub` (truncated subtraction), nnreal and ennreal powers, and measures for `gcongr`.
1 parent 2737bd9 commit 5fcd552

File tree

5 files changed

+13
-13
lines changed

5 files changed

+13
-13
lines changed

Mathlib/Algebra/Order/Sub/Defs.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -103,7 +103,7 @@ theorem add_tsub_le_left : a + b - a ≤ b :=
103103
tsub_le_iff_left.mpr le_rfl
104104
#align add_tsub_le_left add_tsub_le_left
105105

106-
theorem tsub_le_tsub_right (h : a ≤ b) (c : α) : a - c ≤ b - c :=
106+
@[gcongr] theorem tsub_le_tsub_right (h : a ≤ b) (c : α) : a - c ≤ b - c :=
107107
tsub_le_iff_left.mpr <| h.trans le_add_tsub
108108
#align tsub_le_tsub_right tsub_le_tsub_right
109109

@@ -119,11 +119,11 @@ section Cov
119119

120120
variable [CovariantClass α α (· + ·) (· ≤ ·)]
121121

122-
theorem tsub_le_tsub_left (h : a ≤ b) (c : α) : c - b ≤ c - a :=
122+
@[gcongr] theorem tsub_le_tsub_left (h : a ≤ b) (c : α) : c - b ≤ c - a :=
123123
tsub_le_iff_left.mpr <| le_add_tsub.trans <| add_le_add_right h _
124124
#align tsub_le_tsub_left tsub_le_tsub_left
125125

126-
theorem tsub_le_tsub (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c :=
126+
@[gcongr] theorem tsub_le_tsub (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c :=
127127
(tsub_le_tsub_right hab _).trans <| tsub_le_tsub_left hcd _
128128
#align tsub_le_tsub tsub_le_tsub
129129

Mathlib/Analysis/SpecialFunctions/Pow/NNReal.lean

Lines changed: 7 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -198,11 +198,11 @@ theorem _root_.Real.finset_prod_rpow
198198

199199
end Real
200200

201-
theorem rpow_le_rpow {x y : ℝ≥0} {z : ℝ} (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x ^ z ≤ y ^ z :=
201+
@[gcongr] theorem rpow_le_rpow {x y : ℝ≥0} {z : ℝ} (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x ^ z ≤ y ^ z :=
202202
Real.rpow_le_rpow x.2 h₁ h₂
203203
#align nnreal.rpow_le_rpow NNReal.rpow_le_rpow
204204

205-
theorem rpow_lt_rpow {x y : ℝ≥0} {z : ℝ} (h₁ : x < y) (h₂ : 0 < z) : x ^ z < y ^ z :=
205+
@[gcongr] theorem rpow_lt_rpow {x y : ℝ≥0} {z : ℝ} (h₁ : x < y) (h₂ : 0 < z) : x ^ z < y ^ z :=
206206
Real.rpow_lt_rpow x.2 h₁ h₂
207207
#align nnreal.rpow_lt_rpow NNReal.rpow_lt_rpow
208208

@@ -222,12 +222,12 @@ theorem rpow_one_div_le_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ^ (1 / z)
222222
rw [← rpow_le_rpow_iff hz, rpow_self_rpow_inv hz.ne']
223223
#align nnreal.rpow_one_div_le_iff NNReal.rpow_one_div_le_iff
224224

225-
theorem rpow_lt_rpow_of_exponent_lt {x : ℝ≥0} {y z : ℝ} (hx : 1 < x) (hyz : y < z) :
225+
@[gcongr] theorem rpow_lt_rpow_of_exponent_lt {x : ℝ≥0} {y z : ℝ} (hx : 1 < x) (hyz : y < z) :
226226
x ^ y < x ^ z :=
227227
Real.rpow_lt_rpow_of_exponent_lt hx hyz
228228
#align nnreal.rpow_lt_rpow_of_exponent_lt NNReal.rpow_lt_rpow_of_exponent_lt
229229

230-
theorem rpow_le_rpow_of_exponent_le {x : ℝ≥0} {y z : ℝ} (hx : 1 ≤ x) (hyz : y ≤ z) :
230+
@[gcongr] theorem rpow_le_rpow_of_exponent_le {x : ℝ≥0} {y z : ℝ} (hx : 1 ≤ x) (hyz : y ≤ z) :
231231
x ^ y ≤ x ^ z :=
232232
Real.rpow_le_rpow_of_exponent_le hx hyz
233233
#align nnreal.rpow_le_rpow_of_exponent_le NNReal.rpow_le_rpow_of_exponent_le
@@ -639,11 +639,11 @@ theorem orderIsoRpow_symm_apply (y : ℝ) (hy : 0 < y) :
639639
rfl
640640
#align ennreal.order_iso_rpow_symm_apply ENNReal.orderIsoRpow_symm_apply
641641

642-
theorem rpow_le_rpow {x y : ℝ≥0∞} {z : ℝ} (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x ^ z ≤ y ^ z :=
642+
@[gcongr] theorem rpow_le_rpow {x y : ℝ≥0∞} {z : ℝ} (h₁ : x ≤ y) (h₂ : 0 ≤ z) : x ^ z ≤ y ^ z :=
643643
monotone_rpow_of_nonneg h₂ h₁
644644
#align ennreal.rpow_le_rpow ENNReal.rpow_le_rpow
645645

646-
theorem rpow_lt_rpow {x y : ℝ≥0∞} {z : ℝ} (h₁ : x < y) (h₂ : 0 < z) : x ^ z < y ^ z :=
646+
@[gcongr] theorem rpow_lt_rpow {x y : ℝ≥0∞} {z : ℝ} (h₁ : x < y) (h₂ : 0 < z) : x ^ z < y ^ z :=
647647
strictMono_rpow_of_pos h₂ h₁
648648
#align ennreal.rpow_lt_rpow ENNReal.rpow_lt_rpow
649649

@@ -681,7 +681,7 @@ theorem rpow_lt_rpow_of_exponent_lt {x : ℝ≥0∞} {y z : ℝ} (hx : 1 < x) (h
681681
NNReal.rpow_lt_rpow_of_exponent_lt hx hyz]
682682
#align ennreal.rpow_lt_rpow_of_exponent_lt ENNReal.rpow_lt_rpow_of_exponent_lt
683683

684-
theorem rpow_le_rpow_of_exponent_le {x : ℝ≥0∞} {y z : ℝ} (hx : 1 ≤ x) (hyz : y ≤ z) :
684+
@[gcongr] theorem rpow_le_rpow_of_exponent_le {x : ℝ≥0∞} {y z : ℝ} (hx : 1 ≤ x) (hyz : y ≤ z) :
685685
x ^ y ≤ x ^ z := by
686686
cases x
687687
· rcases lt_trichotomy y 0 with (Hy | Hy | Hy) <;>

Mathlib/MeasureTheory/Integral/Layercake.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -237,7 +237,7 @@ theorem lintegral_comp_eq_lintegral_meas_le_mul_of_measurable (μ : Measure α)
237237
apply set_lintegral_mono' measurableSet_Ioc (fun x hx ↦ ?_)
238238
rw [← h's]
239239
gcongr
240-
exact measure_mono (fun a ha ↦ hx.2.trans (le_of_lt ha))
240+
exact fun a ha ↦ hx.2.trans (le_of_lt ha)
241241
_ ≤ ∫⁻ t in Ioi 0, μ {a : α | t ≤ f a} * ENNReal.ofReal (g t) :=
242242
lintegral_mono_set Ioc_subset_Ioi_self
243243
/- The second integral is infinite, as one integrates amont other things on those `ω` where

Mathlib/MeasureTheory/Integral/Lebesgue.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -813,7 +813,7 @@ theorem lintegral_add_mul_meas_add_le_le_lintegral {f g : α → ℝ≥0∞} (hl
813813
by rw [hφ_eq]
814814
_ ≤ ∫⁻ x, φ x ∂μ + ε * μ { x | φ x + ε ≤ g x } := by
815815
gcongr
816-
exact measure_mono fun x => (add_le_add_right (hφ_le _) _).trans
816+
exact fun x => (add_le_add_right (hφ_le _) _).trans
817817
_ = ∫⁻ x, φ x + indicator { x | φ x + ε ≤ g x } (fun _ => ε) x ∂μ := by
818818
rw [lintegral_add_left hφm, lintegral_indicator₀, set_lintegral_const]
819819
exact measurableSet_le (hφm.nullMeasurable.measurable'.add_const _) hg.nullMeasurable

Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -190,7 +190,7 @@ theorem nonempty_of_measure_ne_zero (h : μ s ≠ 0) : s.Nonempty :=
190190
nonempty_iff_ne_empty.2 fun h' => h <| h'.symm ▸ measure_empty
191191
#align measure_theory.nonempty_of_measure_ne_zero MeasureTheory.nonempty_of_measure_ne_zero
192192

193-
theorem measure_mono (h : s₁ ⊆ s₂) : μ s₁ ≤ μ s₂ :=
193+
@[gcongr] theorem measure_mono (h : s₁ ⊆ s₂) : μ s₁ ≤ μ s₂ :=
194194
μ.mono h
195195
#align measure_theory.measure_mono MeasureTheory.measure_mono
196196

0 commit comments

Comments
 (0)