MPE Home Metamath Proof Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  MPE Home  >  Th. List  >  nn0gcdsq Structured version   Visualization version   GIF version

Theorem nn0gcdsq 16554
Description: Squaring commutes with GCD, in particular two coprime numbers have coprime squares. (Contributed by Stefan O'Rear, 15-Sep-2014.)
Assertion
Ref Expression
nn0gcdsq ((𝐴 ∈ ℕ0𝐵 ∈ ℕ0) → ((𝐴 gcd 𝐵)↑2) = ((𝐴↑2) gcd (𝐵↑2)))

Proof of Theorem nn0gcdsq
StepHypRef Expression
1 elnn0 12337 . 2 (𝐴 ∈ ℕ0 ↔ (𝐴 ∈ ℕ ∨ 𝐴 = 0))
2 elnn0 12337 . 2 (𝐵 ∈ ℕ0 ↔ (𝐵 ∈ ℕ ∨ 𝐵 = 0))
3 sqgcd 16368 . . 3 ((𝐴 ∈ ℕ ∧ 𝐵 ∈ ℕ) → ((𝐴 gcd 𝐵)↑2) = ((𝐴↑2) gcd (𝐵↑2)))
4 nncn 12083 . . . . . . 7 (𝐵 ∈ ℕ → 𝐵 ∈ ℂ)
5 abssq 15118 . . . . . . 7 (𝐵 ∈ ℂ → ((abs‘𝐵)↑2) = (abs‘(𝐵↑2)))
64, 5syl 17 . . . . . 6 (𝐵 ∈ ℕ → ((abs‘𝐵)↑2) = (abs‘(𝐵↑2)))
7 nnz 12444 . . . . . . . 8 (𝐵 ∈ ℕ → 𝐵 ∈ ℤ)
8 gcd0id 16326 . . . . . . . 8 (𝐵 ∈ ℤ → (0 gcd 𝐵) = (abs‘𝐵))
97, 8syl 17 . . . . . . 7 (𝐵 ∈ ℕ → (0 gcd 𝐵) = (abs‘𝐵))
109oveq1d 7353 . . . . . 6 (𝐵 ∈ ℕ → ((0 gcd 𝐵)↑2) = ((abs‘𝐵)↑2))
11 sq0 14011 . . . . . . . . 9 (0↑2) = 0
1211a1i 11 . . . . . . . 8 (𝐵 ∈ ℕ → (0↑2) = 0)
1312oveq1d 7353 . . . . . . 7 (𝐵 ∈ ℕ → ((0↑2) gcd (𝐵↑2)) = (0 gcd (𝐵↑2)))
14 zsqcl 13950 . . . . . . . 8 (𝐵 ∈ ℤ → (𝐵↑2) ∈ ℤ)
15 gcd0id 16326 . . . . . . . 8 ((𝐵↑2) ∈ ℤ → (0 gcd (𝐵↑2)) = (abs‘(𝐵↑2)))
167, 14, 153syl 18 . . . . . . 7 (𝐵 ∈ ℕ → (0 gcd (𝐵↑2)) = (abs‘(𝐵↑2)))
1713, 16eqtrd 2776 . . . . . 6 (𝐵 ∈ ℕ → ((0↑2) gcd (𝐵↑2)) = (abs‘(𝐵↑2)))
186, 10, 173eqtr4d 2786 . . . . 5 (𝐵 ∈ ℕ → ((0 gcd 𝐵)↑2) = ((0↑2) gcd (𝐵↑2)))
1918adantl 482 . . . 4 ((𝐴 = 0 ∧ 𝐵 ∈ ℕ) → ((0 gcd 𝐵)↑2) = ((0↑2) gcd (𝐵↑2)))
20 oveq1 7345 . . . . . . 7 (𝐴 = 0 → (𝐴 gcd 𝐵) = (0 gcd 𝐵))
2120oveq1d 7353 . . . . . 6 (𝐴 = 0 → ((𝐴 gcd 𝐵)↑2) = ((0 gcd 𝐵)↑2))
22 oveq1 7345 . . . . . . 7 (𝐴 = 0 → (𝐴↑2) = (0↑2))
2322oveq1d 7353 . . . . . 6 (𝐴 = 0 → ((𝐴↑2) gcd (𝐵↑2)) = ((0↑2) gcd (𝐵↑2)))
2421, 23eqeq12d 2752 . . . . 5 (𝐴 = 0 → (((𝐴 gcd 𝐵)↑2) = ((𝐴↑2) gcd (𝐵↑2)) ↔ ((0 gcd 𝐵)↑2) = ((0↑2) gcd (𝐵↑2))))
2524adantr 481 . . . 4 ((𝐴 = 0 ∧ 𝐵 ∈ ℕ) → (((𝐴 gcd 𝐵)↑2) = ((𝐴↑2) gcd (𝐵↑2)) ↔ ((0 gcd 𝐵)↑2) = ((0↑2) gcd (𝐵↑2))))
2619, 25mpbird 256 . . 3 ((𝐴 = 0 ∧ 𝐵 ∈ ℕ) → ((𝐴 gcd 𝐵)↑2) = ((𝐴↑2) gcd (𝐵↑2)))
27 nncn 12083 . . . . . . 7 (𝐴 ∈ ℕ → 𝐴 ∈ ℂ)
28 abssq 15118 . . . . . . 7 (𝐴 ∈ ℂ → ((abs‘𝐴)↑2) = (abs‘(𝐴↑2)))
2927, 28syl 17 . . . . . 6 (𝐴 ∈ ℕ → ((abs‘𝐴)↑2) = (abs‘(𝐴↑2)))
30 nnz 12444 . . . . . . . 8 (𝐴 ∈ ℕ → 𝐴 ∈ ℤ)
31 gcdid0 16327 . . . . . . . 8 (𝐴 ∈ ℤ → (𝐴 gcd 0) = (abs‘𝐴))
3230, 31syl 17 . . . . . . 7 (𝐴 ∈ ℕ → (𝐴 gcd 0) = (abs‘𝐴))
3332oveq1d 7353 . . . . . 6 (𝐴 ∈ ℕ → ((𝐴 gcd 0)↑2) = ((abs‘𝐴)↑2))
3411a1i 11 . . . . . . . 8 (𝐴 ∈ ℕ → (0↑2) = 0)
3534oveq2d 7354 . . . . . . 7 (𝐴 ∈ ℕ → ((𝐴↑2) gcd (0↑2)) = ((𝐴↑2) gcd 0))
36 zsqcl 13950 . . . . . . . 8 (𝐴 ∈ ℤ → (𝐴↑2) ∈ ℤ)
37 gcdid0 16327 . . . . . . . 8 ((𝐴↑2) ∈ ℤ → ((𝐴↑2) gcd 0) = (abs‘(𝐴↑2)))
3830, 36, 373syl 18 . . . . . . 7 (𝐴 ∈ ℕ → ((𝐴↑2) gcd 0) = (abs‘(𝐴↑2)))
3935, 38eqtrd 2776 . . . . . 6 (𝐴 ∈ ℕ → ((𝐴↑2) gcd (0↑2)) = (abs‘(𝐴↑2)))
4029, 33, 393eqtr4d 2786 . . . . 5 (𝐴 ∈ ℕ → ((𝐴 gcd 0)↑2) = ((𝐴↑2) gcd (0↑2)))
4140adantr 481 . . . 4 ((𝐴 ∈ ℕ ∧ 𝐵 = 0) → ((𝐴 gcd 0)↑2) = ((𝐴↑2) gcd (0↑2)))
42 oveq2 7346 . . . . . . 7 (𝐵 = 0 → (𝐴 gcd 𝐵) = (𝐴 gcd 0))
4342oveq1d 7353 . . . . . 6 (𝐵 = 0 → ((𝐴 gcd 𝐵)↑2) = ((𝐴 gcd 0)↑2))
44 oveq1 7345 . . . . . . 7 (𝐵 = 0 → (𝐵↑2) = (0↑2))
4544oveq2d 7354 . . . . . 6 (𝐵 = 0 → ((𝐴↑2) gcd (𝐵↑2)) = ((𝐴↑2) gcd (0↑2)))
4643, 45eqeq12d 2752 . . . . 5 (𝐵 = 0 → (((𝐴 gcd 𝐵)↑2) = ((𝐴↑2) gcd (𝐵↑2)) ↔ ((𝐴 gcd 0)↑2) = ((𝐴↑2) gcd (0↑2))))
4746adantl 482 . . . 4 ((𝐴 ∈ ℕ ∧ 𝐵 = 0) → (((𝐴 gcd 𝐵)↑2) = ((𝐴↑2) gcd (𝐵↑2)) ↔ ((𝐴 gcd 0)↑2) = ((𝐴↑2) gcd (0↑2))))
4841, 47mpbird 256 . . 3 ((𝐴 ∈ ℕ ∧ 𝐵 = 0) → ((𝐴 gcd 𝐵)↑2) = ((𝐴↑2) gcd (𝐵↑2)))
49 gcd0val 16304 . . . . . 6 (0 gcd 0) = 0
5049oveq1i 7348 . . . . 5 ((0 gcd 0)↑2) = (0↑2)
5111, 11oveq12i 7350 . . . . . 6 ((0↑2) gcd (0↑2)) = (0 gcd 0)
5251, 49eqtri 2764 . . . . 5 ((0↑2) gcd (0↑2)) = 0
5311, 50, 523eqtr4i 2774 . . . 4 ((0 gcd 0)↑2) = ((0↑2) gcd (0↑2))
54 oveq12 7347 . . . . 5 ((𝐴 = 0 ∧ 𝐵 = 0) → (𝐴 gcd 𝐵) = (0 gcd 0))
5554oveq1d 7353 . . . 4 ((𝐴 = 0 ∧ 𝐵 = 0) → ((𝐴 gcd 𝐵)↑2) = ((0 gcd 0)↑2))
5622, 44oveqan12d 7357 . . . 4 ((𝐴 = 0 ∧ 𝐵 = 0) → ((𝐴↑2) gcd (𝐵↑2)) = ((0↑2) gcd (0↑2)))
5753, 55, 563eqtr4a 2802 . . 3 ((𝐴 = 0 ∧ 𝐵 = 0) → ((𝐴 gcd 𝐵)↑2) = ((𝐴↑2) gcd (𝐵↑2)))
583, 26, 48, 57ccase 1035 . 2 (((𝐴 ∈ ℕ ∨ 𝐴 = 0) ∧ (𝐵 ∈ ℕ ∨ 𝐵 = 0)) → ((𝐴 gcd 𝐵)↑2) = ((𝐴↑2) gcd (𝐵↑2)))
591, 2, 58syl2anb 598 1 ((𝐴 ∈ ℕ0𝐵 ∈ ℕ0) → ((𝐴 gcd 𝐵)↑2) = ((𝐴↑2) gcd (𝐵↑2)))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 205  wa 396  wo 844   = wceq 1540  wcel 2105  cfv 6480  (class class class)co 7338  cc 10971  0cc0 10973  cn 12075  2c2 12130  0cn0 12335  cz 12421  cexp 13884  abscabs 15045   gcd cgcd 16301
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1796  ax-4 1810  ax-5 1912  ax-6 1970  ax-7 2010  ax-8 2107  ax-9 2115  ax-10 2136  ax-11 2153  ax-12 2170  ax-ext 2707  ax-sep 5244  ax-nul 5251  ax-pow 5309  ax-pr 5373  ax-un 7651  ax-cnex 11029  ax-resscn 11030  ax-1cn 11031  ax-icn 11032  ax-addcl 11033  ax-addrcl 11034  ax-mulcl 11035  ax-mulrcl 11036  ax-mulcom 11037  ax-addass 11038  ax-mulass 11039  ax-distr 11040  ax-i2m1 11041  ax-1ne0 11042  ax-1rid 11043  ax-rnegex 11044  ax-rrecex 11045  ax-cnre 11046  ax-pre-lttri 11047  ax-pre-lttrn 11048  ax-pre-ltadd 11049  ax-pre-mulgt0 11050  ax-pre-sup 11051
This theorem depends on definitions:  df-bi 206  df-an 397  df-or 845  df-3or 1087  df-3an 1088  df-tru 1543  df-fal 1553  df-ex 1781  df-nf 1785  df-sb 2067  df-mo 2538  df-eu 2567  df-clab 2714  df-cleq 2728  df-clel 2814  df-nfc 2886  df-ne 2941  df-nel 3047  df-ral 3062  df-rex 3071  df-rmo 3349  df-reu 3350  df-rab 3404  df-v 3443  df-sbc 3728  df-csb 3844  df-dif 3901  df-un 3903  df-in 3905  df-ss 3915  df-pss 3917  df-nul 4271  df-if 4475  df-pw 4550  df-sn 4575  df-pr 4577  df-op 4581  df-uni 4854  df-iun 4944  df-br 5094  df-opab 5156  df-mpt 5177  df-tr 5211  df-id 5519  df-eprel 5525  df-po 5533  df-so 5534  df-fr 5576  df-we 5578  df-xp 5627  df-rel 5628  df-cnv 5629  df-co 5630  df-dm 5631  df-rn 5632  df-res 5633  df-ima 5634  df-pred 6239  df-ord 6306  df-on 6307  df-lim 6308  df-suc 6309  df-iota 6432  df-fun 6482  df-fn 6483  df-f 6484  df-f1 6485  df-fo 6486  df-f1o 6487  df-fv 6488  df-riota 7294  df-ov 7341  df-oprab 7342  df-mpo 7343  df-om 7782  df-2nd 7901  df-frecs 8168  df-wrecs 8199  df-recs 8273  df-rdg 8312  df-er 8570  df-en 8806  df-dom 8807  df-sdom 8808  df-sup 9300  df-inf 9301  df-pnf 11113  df-mnf 11114  df-xr 11115  df-ltxr 11116  df-le 11117  df-sub 11309  df-neg 11310  df-div 11735  df-nn 12076  df-2 12138  df-3 12139  df-n0 12336  df-z 12422  df-uz 12685  df-rp 12833  df-fl 13614  df-mod 13692  df-seq 13824  df-exp 13885  df-cj 14910  df-re 14911  df-im 14912  df-sqrt 15046  df-abs 15047  df-dvds 16064  df-gcd 16302
This theorem is referenced by:  zgcdsq  16555
  Copyright terms: Public domain W3C validator