@@ -198,11 +198,11 @@ theorem _root_.Real.finset_prod_rpow
198
198
199
199
end Real
200
200
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 :=
202
202
Real.rpow_le_rpow x.2 h₁ h₂
203
203
#align nnreal.rpow_le_rpow NNReal.rpow_le_rpow
204
204
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 :=
206
206
Real.rpow_lt_rpow x.2 h₁ h₂
207
207
#align nnreal.rpow_lt_rpow NNReal.rpow_lt_rpow
208
208
@@ -222,12 +222,12 @@ theorem rpow_one_div_le_iff {x y : ℝ≥0} {z : ℝ} (hz : 0 < z) : x ^ (1 / z)
222
222
rw [← rpow_le_rpow_iff hz, rpow_self_rpow_inv hz.ne']
223
223
#align nnreal.rpow_one_div_le_iff NNReal.rpow_one_div_le_iff
224
224
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) :
226
226
x ^ y < x ^ z :=
227
227
Real.rpow_lt_rpow_of_exponent_lt hx hyz
228
228
#align nnreal.rpow_lt_rpow_of_exponent_lt NNReal.rpow_lt_rpow_of_exponent_lt
229
229
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) :
231
231
x ^ y ≤ x ^ z :=
232
232
Real.rpow_le_rpow_of_exponent_le hx hyz
233
233
#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) :
639
639
rfl
640
640
#align ennreal.order_iso_rpow_symm_apply ENNReal.orderIsoRpow_symm_apply
641
641
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 :=
643
643
monotone_rpow_of_nonneg h₂ h₁
644
644
#align ennreal.rpow_le_rpow ENNReal.rpow_le_rpow
645
645
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 :=
647
647
strictMono_rpow_of_pos h₂ h₁
648
648
#align ennreal.rpow_lt_rpow ENNReal.rpow_lt_rpow
649
649
@@ -681,7 +681,7 @@ theorem rpow_lt_rpow_of_exponent_lt {x : ℝ≥0∞} {y z : ℝ} (hx : 1 < x) (h
681
681
NNReal.rpow_lt_rpow_of_exponent_lt hx hyz]
682
682
#align ennreal.rpow_lt_rpow_of_exponent_lt ENNReal.rpow_lt_rpow_of_exponent_lt
683
683
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) :
685
685
x ^ y ≤ x ^ z := by
686
686
cases x
687
687
· rcases lt_trichotomy y 0 with (Hy | Hy | Hy) <;>
0 commit comments