Metamath Proof Explorer


Theorem sqrecd

Description: Square of reciprocal. (Contributed by Mario Carneiro, 28-May-2016)

Ref Expression
Hypotheses expcld.1 ( 𝜑𝐴 ∈ ℂ )
sqrecd.1 ( 𝜑𝐴 ≠ 0 )
Assertion sqrecd ( 𝜑 → ( ( 1 / 𝐴 ) ↑ 2 ) = ( 1 / ( 𝐴 ↑ 2 ) ) )

Proof

Step Hyp Ref Expression
1 expcld.1 ( 𝜑𝐴 ∈ ℂ )
2 sqrecd.1 ( 𝜑𝐴 ≠ 0 )
3 2z 2 ∈ ℤ
4 3 a1i ( 𝜑 → 2 ∈ ℤ )
5 exprec ( ( 𝐴 ∈ ℂ ∧ 𝐴 ≠ 0 ∧ 2 ∈ ℤ ) → ( ( 1 / 𝐴 ) ↑ 2 ) = ( 1 / ( 𝐴 ↑ 2 ) ) )
6 1 2 4 5 syl3anc ( 𝜑 → ( ( 1 / 𝐴 ) ↑ 2 ) = ( 1 / ( 𝐴 ↑ 2 ) ) )