![]() |
Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
|
Mirrors > Home > MPE Home > Th. List > lgsquad | Structured version Visualization version GIF version |
Description: The Law of Quadratic Reciprocity, see also theorem 9.8 in [ApostolNT] p. 185. If 𝑃 and 𝑄 are distinct odd primes, then the product of the Legendre symbols (𝑃 /L 𝑄) and (𝑄 /L 𝑃) is the parity of ((𝑃 − 1) / 2) · ((𝑄 − 1) / 2). This uses Eisenstein's proof, which also has a nice geometric interpretation - see https://en.wikipedia.org/wiki/Proofs_of_quadratic_reciprocity. This is Metamath 100 proof #7. (Contributed by Mario Carneiro, 19-Jun-2015.) |
Ref | Expression |
---|---|
lgsquad | ⊢ ((𝑃 ∈ (ℙ ∖ {2}) ∧ 𝑄 ∈ (ℙ ∖ {2}) ∧ 𝑃 ≠ 𝑄) → ((𝑃 /L 𝑄) · (𝑄 /L 𝑃)) = (-1↑(((𝑃 − 1) / 2) · ((𝑄 − 1) / 2)))) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | simp1 1127 | . 2 ⊢ ((𝑃 ∈ (ℙ ∖ {2}) ∧ 𝑄 ∈ (ℙ ∖ {2}) ∧ 𝑃 ≠ 𝑄) → 𝑃 ∈ (ℙ ∖ {2})) | |
2 | simp2 1128 | . 2 ⊢ ((𝑃 ∈ (ℙ ∖ {2}) ∧ 𝑄 ∈ (ℙ ∖ {2}) ∧ 𝑃 ≠ 𝑄) → 𝑄 ∈ (ℙ ∖ {2})) | |
3 | simp3 1129 | . 2 ⊢ ((𝑃 ∈ (ℙ ∖ {2}) ∧ 𝑄 ∈ (ℙ ∖ {2}) ∧ 𝑃 ≠ 𝑄) → 𝑃 ≠ 𝑄) | |
4 | eqid 2777 | . 2 ⊢ ((𝑃 − 1) / 2) = ((𝑃 − 1) / 2) | |
5 | eqid 2777 | . 2 ⊢ ((𝑄 − 1) / 2) = ((𝑄 − 1) / 2) | |
6 | eleq1w 2841 | . . . . 5 ⊢ (𝑥 = 𝑧 → (𝑥 ∈ (1...((𝑃 − 1) / 2)) ↔ 𝑧 ∈ (1...((𝑃 − 1) / 2)))) | |
7 | eleq1w 2841 | . . . . 5 ⊢ (𝑦 = 𝑤 → (𝑦 ∈ (1...((𝑄 − 1) / 2)) ↔ 𝑤 ∈ (1...((𝑄 − 1) / 2)))) | |
8 | 6, 7 | bi2anan9 629 | . . . 4 ⊢ ((𝑥 = 𝑧 ∧ 𝑦 = 𝑤) → ((𝑥 ∈ (1...((𝑃 − 1) / 2)) ∧ 𝑦 ∈ (1...((𝑄 − 1) / 2))) ↔ (𝑧 ∈ (1...((𝑃 − 1) / 2)) ∧ 𝑤 ∈ (1...((𝑄 − 1) / 2))))) |
9 | oveq1 6929 | . . . . 5 ⊢ (𝑦 = 𝑤 → (𝑦 · 𝑃) = (𝑤 · 𝑃)) | |
10 | oveq1 6929 | . . . . 5 ⊢ (𝑥 = 𝑧 → (𝑥 · 𝑄) = (𝑧 · 𝑄)) | |
11 | 9, 10 | breqan12rd 4903 | . . . 4 ⊢ ((𝑥 = 𝑧 ∧ 𝑦 = 𝑤) → ((𝑦 · 𝑃) < (𝑥 · 𝑄) ↔ (𝑤 · 𝑃) < (𝑧 · 𝑄))) |
12 | 8, 11 | anbi12d 624 | . . 3 ⊢ ((𝑥 = 𝑧 ∧ 𝑦 = 𝑤) → (((𝑥 ∈ (1...((𝑃 − 1) / 2)) ∧ 𝑦 ∈ (1...((𝑄 − 1) / 2))) ∧ (𝑦 · 𝑃) < (𝑥 · 𝑄)) ↔ ((𝑧 ∈ (1...((𝑃 − 1) / 2)) ∧ 𝑤 ∈ (1...((𝑄 − 1) / 2))) ∧ (𝑤 · 𝑃) < (𝑧 · 𝑄)))) |
13 | 12 | cbvopabv 4958 | . 2 ⊢ {〈𝑥, 𝑦〉 ∣ ((𝑥 ∈ (1...((𝑃 − 1) / 2)) ∧ 𝑦 ∈ (1...((𝑄 − 1) / 2))) ∧ (𝑦 · 𝑃) < (𝑥 · 𝑄))} = {〈𝑧, 𝑤〉 ∣ ((𝑧 ∈ (1...((𝑃 − 1) / 2)) ∧ 𝑤 ∈ (1...((𝑄 − 1) / 2))) ∧ (𝑤 · 𝑃) < (𝑧 · 𝑄))} |
14 | 1, 2, 3, 4, 5, 13 | lgsquadlem3 25559 | 1 ⊢ ((𝑃 ∈ (ℙ ∖ {2}) ∧ 𝑄 ∈ (ℙ ∖ {2}) ∧ 𝑃 ≠ 𝑄) → ((𝑃 /L 𝑄) · (𝑄 /L 𝑃)) = (-1↑(((𝑃 − 1) / 2) · ((𝑄 − 1) / 2)))) |
Colors of variables: wff setvar class |
Syntax hints: → wi 4 ∧ wa 386 ∧ w3a 1071 = wceq 1601 ∈ wcel 2106 ≠ wne 2968 ∖ cdif 3788 {csn 4397 class class class wbr 4886 {copab 4948 (class class class)co 6922 1c1 10273 · cmul 10277 < clt 10411 − cmin 10606 -cneg 10607 / cdiv 11032 2c2 11430 ...cfz 12643 ↑cexp 13178 ℙcprime 15790 /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 2054 ax-8 2108 ax-9 2115 ax-10 2134 ax-11 2149 ax-12 2162 ax-13 2333 ax-ext 2753 ax-rep 5006 ax-sep 5017 ax-nul 5025 ax-pow 5077 ax-pr 5138 ax-un 7226 ax-inf2 8835 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 ax-addf 10351 ax-mulf 10352 |
This theorem depends on definitions: df-bi 199 df-an 387 df-or 837 df-3or 1072 df-3an 1073 df-tru 1605 df-fal 1615 df-ex 1824 df-nf 1828 df-sb 2012 df-mo 2550 df-eu 2586 df-clab 2763 df-cleq 2769 df-clel 2773 df-nfc 2920 df-ne 2969 df-nel 3075 df-ral 3094 df-rex 3095 df-reu 3096 df-rmo 3097 df-rab 3098 df-v 3399 df-sbc 3652 df-csb 3751 df-dif 3794 df-un 3796 df-in 3798 df-ss 3805 df-pss 3807 df-nul 4141 df-if 4307 df-pw 4380 df-sn 4398 df-pr 4400 df-tp 4402 df-op 4404 df-uni 4672 df-int 4711 df-iun 4755 df-disj 4855 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-se 5315 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-isom 6144 df-riota 6883 df-ov 6925 df-oprab 6926 df-mpt2 6927 df-of 7174 df-om 7344 df-1st 7445 df-2nd 7446 df-supp 7577 df-tpos 7634 df-wrecs 7689 df-recs 7751 df-rdg 7789 df-1o 7843 df-2o 7844 df-oadd 7847 df-er 8026 df-ec 8028 df-qs 8032 df-map 8142 df-en 8242 df-dom 8243 df-sdom 8244 df-fin 8245 df-fsupp 8564 df-sup 8636 df-inf 8637 df-oi 8704 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-4 11440 df-5 11441 df-6 11442 df-7 11443 df-8 11444 df-9 11445 df-n0 11643 df-xnn0 11715 df-z 11729 df-dec 11846 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-clim 14627 df-sum 14825 df-dvds 15388 df-gcd 15623 df-prm 15791 df-phi 15875 df-pc 15946 df-struct 16257 df-ndx 16258 df-slot 16259 df-base 16261 df-sets 16262 df-ress 16263 df-plusg 16351 df-mulr 16352 df-starv 16353 df-sca 16354 df-vsca 16355 df-ip 16356 df-tset 16357 df-ple 16358 df-ds 16360 df-unif 16361 df-0g 16488 df-gsum 16489 df-imas 16554 df-qus 16555 df-mgm 17628 df-sgrp 17670 df-mnd 17681 df-mhm 17721 df-submnd 17722 df-grp 17812 df-minusg 17813 df-sbg 17814 df-mulg 17928 df-subg 17975 df-nsg 17976 df-eqg 17977 df-ghm 18042 df-cntz 18133 df-cmn 18581 df-abl 18582 df-mgp 18877 df-ur 18889 df-ring 18936 df-cring 18937 df-oppr 19010 df-dvdsr 19028 df-unit 19029 df-invr 19059 df-dvr 19070 df-rnghom 19104 df-drng 19141 df-field 19142 df-subrg 19170 df-lmod 19257 df-lss 19325 df-lsp 19367 df-sra 19569 df-rgmod 19570 df-lidl 19571 df-rsp 19572 df-2idl 19629 df-nzr 19655 df-rlreg 19680 df-domn 19681 df-idom 19682 df-cnfld 20143 df-zring 20215 df-zrh 20248 df-zn 20251 df-lgs 25472 |
This theorem is referenced by: lgsquad2 25563 |
Copyright terms: Public domain | W3C validator |