Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit c7202e5

Browse files
feat(analysis/exponential): pow_nat_rpow_nat_inv (#740)
1 parent 78ce6e4 commit c7202e5

File tree

1 file changed

+8
-0
lines changed

1 file changed

+8
-0
lines changed

src/analysis/exponential.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1115,6 +1115,9 @@ begin
11151115
(complex.of_real_mul _ _).symm, -complex.of_real_mul] at *
11161116
end
11171117

1118+
@[simp] lemma abs_cpow_inv_nat (x : ℂ) (n : ℕ) : abs (x ^ (n⁻¹ : ℂ)) = x.abs ^ (n⁻¹ : ℝ) :=
1119+
by rw ← abs_cpow_real; simp [-abs_cpow_real]
1120+
11181121
end complex
11191122

11201123
namespace real
@@ -1152,4 +1155,9 @@ by simp only [rpow_def, (complex.of_real_pow _ _).symm, complex.cpow_nat_cast,
11521155
by simp only [rpow_def, (complex.of_real_fpow _ _).symm, complex.cpow_int_cast,
11531156
complex.of_real_int_cast, complex.of_real_re]
11541157

1158+
lemma pow_nat_rpow_nat_inv {x : ℝ} (hx : 0 ≤ x) {n : ℕ} (hn : 0 < n) :
1159+
(x ^ n) ^ (n⁻¹ : ℝ) = x :=
1160+
have hn0 : (n : ℝ) ≠ 0, by simpa [nat.pos_iff_ne_zero'] using hn,
1161+
by rw [← rpow_nat_cast, ← rpow_mul hx, mul_inv_cancel hn0, rpow_one]
1162+
11551163
end real

0 commit comments

Comments
 (0)