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

Theorem lgssq2 25515
Description: The Legendre symbol at a square is equal to 1. (Contributed by Mario Carneiro, 5-Feb-2015.)
Assertion
Ref Expression
lgssq2 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → (𝐴 /L (𝑁↑2)) = 1)

Proof of Theorem lgssq2
StepHypRef Expression
1 simp1 1127 . . 3 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → 𝐴 ∈ ℤ)
2 nnz 11751 . . . 4 (𝑁 ∈ ℕ → 𝑁 ∈ ℤ)
323ad2ant2 1125 . . 3 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → 𝑁 ∈ ℤ)
4 nnne0 11410 . . . 4 (𝑁 ∈ ℕ → 𝑁 ≠ 0)
543ad2ant2 1125 . . 3 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → 𝑁 ≠ 0)
6 lgsdi 25511 . . 3 (((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝑁 ≠ 0 ∧ 𝑁 ≠ 0)) → (𝐴 /L (𝑁 · 𝑁)) = ((𝐴 /L 𝑁) · (𝐴 /L 𝑁)))
71, 3, 3, 5, 5, 6syl32anc 1446 . 2 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → (𝐴 /L (𝑁 · 𝑁)) = ((𝐴 /L 𝑁) · (𝐴 /L 𝑁)))
8 nncn 11383 . . . . 5 (𝑁 ∈ ℕ → 𝑁 ∈ ℂ)
983ad2ant2 1125 . . . 4 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → 𝑁 ∈ ℂ)
109sqvald 13324 . . 3 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → (𝑁↑2) = (𝑁 · 𝑁))
1110oveq2d 6938 . 2 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → (𝐴 /L (𝑁↑2)) = (𝐴 /L (𝑁 · 𝑁)))
12 lgscl 25488 . . . . . 6 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐴 /L 𝑁) ∈ ℤ)
131, 3, 12syl2anc 579 . . . . 5 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → (𝐴 /L 𝑁) ∈ ℤ)
1413zred 11834 . . . 4 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → (𝐴 /L 𝑁) ∈ ℝ)
15 absresq 14449 . . . 4 ((𝐴 /L 𝑁) ∈ ℝ → ((abs‘(𝐴 /L 𝑁))↑2) = ((𝐴 /L 𝑁)↑2))
1614, 15syl 17 . . 3 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → ((abs‘(𝐴 /L 𝑁))↑2) = ((𝐴 /L 𝑁)↑2))
17 lgsabs1 25513 . . . . . . 7 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((abs‘(𝐴 /L 𝑁)) = 1 ↔ (𝐴 gcd 𝑁) = 1))
182, 17sylan2 586 . . . . . 6 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ) → ((abs‘(𝐴 /L 𝑁)) = 1 ↔ (𝐴 gcd 𝑁) = 1))
1918biimp3ar 1543 . . . . 5 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → (abs‘(𝐴 /L 𝑁)) = 1)
2019oveq1d 6937 . . . 4 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → ((abs‘(𝐴 /L 𝑁))↑2) = (1↑2))
21 sq1 13277 . . . 4 (1↑2) = 1
2220, 21syl6eq 2830 . . 3 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → ((abs‘(𝐴 /L 𝑁))↑2) = 1)
2313zcnd 11835 . . . 4 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → (𝐴 /L 𝑁) ∈ ℂ)
2423sqvald 13324 . . 3 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → ((𝐴 /L 𝑁)↑2) = ((𝐴 /L 𝑁) · (𝐴 /L 𝑁)))
2516, 22, 243eqtr3d 2822 . 2 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → 1 = ((𝐴 /L 𝑁) · (𝐴 /L 𝑁)))
267, 11, 253eqtr4d 2824 1 ((𝐴 ∈ ℤ ∧ 𝑁 ∈ ℕ ∧ (𝐴 gcd 𝑁) = 1) → (𝐴 /L (𝑁↑2)) = 1)
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 198  w3a 1071   = wceq 1601  wcel 2107  wne 2969  cfv 6135  (class class class)co 6922  cc 10270  cr 10271  0cc0 10272  1c1 10273   · cmul 10277  cn 11374  2c2 11430  cz 11728  cexp 13178  abscabs 14381   gcd cgcd 15622   /L clgs 25471
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1839  ax-4 1853  ax-5 1953  ax-6 2021  ax-7 2055  ax-8 2109  ax-9 2116  ax-10 2135  ax-11 2150  ax-12 2163  ax-13 2334  ax-ext 2754  ax-rep 5006  ax-sep 5017  ax-nul 5025  ax-pow 5077  ax-pr 5138  ax-un 7226  ax-cnex 10328  ax-resscn 10329  ax-1cn 10330  ax-icn 10331  ax-addcl 10332  ax-addrcl 10333  ax-mulcl 10334  ax-mulrcl 10335  ax-mulcom 10336  ax-addass 10337  ax-mulass 10338  ax-distr 10339  ax-i2m1 10340  ax-1ne0 10341  ax-1rid 10342  ax-rnegex 10343  ax-rrecex 10344  ax-cnre 10345  ax-pre-lttri 10346  ax-pre-lttrn 10347  ax-pre-ltadd 10348  ax-pre-mulgt0 10349  ax-pre-sup 10350
This theorem depends on definitions:  df-bi 199  df-an 387  df-or 837  df-3or 1072  df-3an 1073  df-tru 1605  df-ex 1824  df-nf 1828  df-sb 2012  df-mo 2551  df-eu 2587  df-clab 2764  df-cleq 2770  df-clel 2774  df-nfc 2921  df-ne 2970  df-nel 3076  df-ral 3095  df-rex 3096  df-reu 3097  df-rmo 3098  df-rab 3099  df-v 3400  df-sbc 3653  df-csb 3752  df-dif 3795  df-un 3797  df-in 3799  df-ss 3806  df-pss 3808  df-nul 4142  df-if 4308  df-pw 4381  df-sn 4399  df-pr 4401  df-tp 4403  df-op 4405  df-uni 4672  df-int 4711  df-iun 4755  df-br 4887  df-opab 4949  df-mpt 4966  df-tr 4988  df-id 5261  df-eprel 5266  df-po 5274  df-so 5275  df-fr 5314  df-we 5316  df-xp 5361  df-rel 5362  df-cnv 5363  df-co 5364  df-dm 5365  df-rn 5366  df-res 5367  df-ima 5368  df-pred 5933  df-ord 5979  df-on 5980  df-lim 5981  df-suc 5982  df-iota 6099  df-fun 6137  df-fn 6138  df-f 6139  df-f1 6140  df-fo 6141  df-f1o 6142  df-fv 6143  df-riota 6883  df-ov 6925  df-oprab 6926  df-mpt2 6927  df-om 7344  df-1st 7445  df-2nd 7446  df-wrecs 7689  df-recs 7751  df-rdg 7789  df-1o 7843  df-2o 7844  df-oadd 7847  df-er 8026  df-map 8142  df-en 8242  df-dom 8243  df-sdom 8244  df-fin 8245  df-sup 8636  df-inf 8637  df-card 9098  df-cda 9325  df-pnf 10413  df-mnf 10414  df-xr 10415  df-ltxr 10416  df-le 10417  df-sub 10608  df-neg 10609  df-div 11033  df-nn 11375  df-2 11438  df-3 11439  df-n0 11643  df-xnn0 11715  df-z 11729  df-uz 11993  df-q 12096  df-rp 12138  df-fz 12644  df-fzo 12785  df-fl 12912  df-mod 12988  df-seq 13120  df-exp 13179  df-hash 13436  df-cj 14246  df-re 14247  df-im 14248  df-sqrt 14382  df-abs 14383  df-dvds 15388  df-gcd 15623  df-prm 15791  df-phi 15875  df-pc 15946  df-lgs 25472
This theorem is referenced by:  lgs1  25518  lgsquad2lem2  25562
  Copyright terms: Public domain W3C validator