![]() |
Intuitionistic Logic Explorer |
< Previous
Next >
Nearby theorems |
|
Mirrors > Home > ILE Home > Th. List > lgsquad | 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 999 | . 2 ⊢ ((𝑃 ∈ (ℙ ∖ {2}) ∧ 𝑄 ∈ (ℙ ∖ {2}) ∧ 𝑃 ≠ 𝑄) → 𝑃 ∈ (ℙ ∖ {2})) | |
2 | simp2 1000 | . 2 ⊢ ((𝑃 ∈ (ℙ ∖ {2}) ∧ 𝑄 ∈ (ℙ ∖ {2}) ∧ 𝑃 ≠ 𝑄) → 𝑄 ∈ (ℙ ∖ {2})) | |
3 | simp3 1001 | . 2 ⊢ ((𝑃 ∈ (ℙ ∖ {2}) ∧ 𝑄 ∈ (ℙ ∖ {2}) ∧ 𝑃 ≠ 𝑄) → 𝑃 ≠ 𝑄) | |
4 | eqid 2196 | . 2 ⊢ ((𝑃 − 1) / 2) = ((𝑃 − 1) / 2) | |
5 | eqid 2196 | . 2 ⊢ ((𝑄 − 1) / 2) = ((𝑄 − 1) / 2) | |
6 | eleq1w 2257 | . . . . 5 ⊢ (𝑥 = 𝑧 → (𝑥 ∈ (1...((𝑃 − 1) / 2)) ↔ 𝑧 ∈ (1...((𝑃 − 1) / 2)))) | |
7 | eleq1w 2257 | . . . . 5 ⊢ (𝑦 = 𝑤 → (𝑦 ∈ (1...((𝑄 − 1) / 2)) ↔ 𝑤 ∈ (1...((𝑄 − 1) / 2)))) | |
8 | 6, 7 | bi2anan9 606 | . . . 4 ⊢ ((𝑥 = 𝑧 ∧ 𝑦 = 𝑤) → ((𝑥 ∈ (1...((𝑃 − 1) / 2)) ∧ 𝑦 ∈ (1...((𝑄 − 1) / 2))) ↔ (𝑧 ∈ (1...((𝑃 − 1) / 2)) ∧ 𝑤 ∈ (1...((𝑄 − 1) / 2))))) |
9 | oveq1 5929 | . . . . 5 ⊢ (𝑦 = 𝑤 → (𝑦 · 𝑃) = (𝑤 · 𝑃)) | |
10 | oveq1 5929 | . . . . 5 ⊢ (𝑥 = 𝑧 → (𝑥 · 𝑄) = (𝑧 · 𝑄)) | |
11 | 9, 10 | breqan12rd 4050 | . . . 4 ⊢ ((𝑥 = 𝑧 ∧ 𝑦 = 𝑤) → ((𝑦 · 𝑃) < (𝑥 · 𝑄) ↔ (𝑤 · 𝑃) < (𝑧 · 𝑄))) |
12 | 8, 11 | anbi12d 473 | . . 3 ⊢ ((𝑥 = 𝑧 ∧ 𝑦 = 𝑤) → (((𝑥 ∈ (1...((𝑃 − 1) / 2)) ∧ 𝑦 ∈ (1...((𝑄 − 1) / 2))) ∧ (𝑦 · 𝑃) < (𝑥 · 𝑄)) ↔ ((𝑧 ∈ (1...((𝑃 − 1) / 2)) ∧ 𝑤 ∈ (1...((𝑄 − 1) / 2))) ∧ (𝑤 · 𝑃) < (𝑧 · 𝑄)))) |
13 | 12 | cbvopabv 4105 | . 2 ⊢ {〈𝑥, 𝑦〉 ∣ ((𝑥 ∈ (1...((𝑃 − 1) / 2)) ∧ 𝑦 ∈ (1...((𝑄 − 1) / 2))) ∧ (𝑦 · 𝑃) < (𝑥 · 𝑄))} = {〈𝑧, 𝑤〉 ∣ ((𝑧 ∈ (1...((𝑃 − 1) / 2)) ∧ 𝑤 ∈ (1...((𝑄 − 1) / 2))) ∧ (𝑤 · 𝑃) < (𝑧 · 𝑄))} |
14 | 1, 2, 3, 4, 5, 13 | lgsquadlem3 15287 | 1 ⊢ ((𝑃 ∈ (ℙ ∖ {2}) ∧ 𝑄 ∈ (ℙ ∖ {2}) ∧ 𝑃 ≠ 𝑄) → ((𝑃 /L 𝑄) · (𝑄 /L 𝑃)) = (-1↑(((𝑃 − 1) / 2) · ((𝑄 − 1) / 2)))) |
Colors of variables: wff set class |
Syntax hints: → wi 4 ∧ wa 104 ∧ w3a 980 = wceq 1364 ∈ wcel 2167 ≠ wne 2367 ∖ cdif 3154 {csn 3622 class class class wbr 4033 {copab 4093 (class class class)co 5922 1c1 7878 · cmul 7882 < clt 8059 − cmin 8195 -cneg 8196 / cdiv 8696 2c2 9038 ...cfz 10080 ↑cexp 10615 ℙcprime 12251 /L clgs 15205 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-ia1 106 ax-ia2 107 ax-ia3 108 ax-in1 615 ax-in2 616 ax-io 710 ax-5 1461 ax-7 1462 ax-gen 1463 ax-ie1 1507 ax-ie2 1508 ax-8 1518 ax-10 1519 ax-11 1520 ax-i12 1521 ax-bndl 1523 ax-4 1524 ax-17 1540 ax-i9 1544 ax-ial 1548 ax-i5r 1549 ax-13 2169 ax-14 2170 ax-ext 2178 ax-coll 4148 ax-sep 4151 ax-nul 4159 ax-pow 4207 ax-pr 4242 ax-un 4468 ax-setind 4573 ax-iinf 4624 ax-cnex 7968 ax-resscn 7969 ax-1cn 7970 ax-1re 7971 ax-icn 7972 ax-addcl 7973 ax-addrcl 7974 ax-mulcl 7975 ax-mulrcl 7976 ax-addcom 7977 ax-mulcom 7978 ax-addass 7979 ax-mulass 7980 ax-distr 7981 ax-i2m1 7982 ax-0lt1 7983 ax-1rid 7984 ax-0id 7985 ax-rnegex 7986 ax-precex 7987 ax-cnre 7988 ax-pre-ltirr 7989 ax-pre-ltwlin 7990 ax-pre-lttrn 7991 ax-pre-apti 7992 ax-pre-ltadd 7993 ax-pre-mulgt0 7994 ax-pre-mulext 7995 ax-arch 7996 ax-caucvg 7997 ax-addf 7999 ax-mulf 8000 |
This theorem depends on definitions: df-bi 117 df-stab 832 df-dc 836 df-3or 981 df-3an 982 df-tru 1367 df-fal 1370 df-xor 1387 df-nf 1475 df-sb 1777 df-eu 2048 df-mo 2049 df-clab 2183 df-cleq 2189 df-clel 2192 df-nfc 2328 df-ne 2368 df-nel 2463 df-ral 2480 df-rex 2481 df-reu 2482 df-rmo 2483 df-rab 2484 df-v 2765 df-sbc 2990 df-csb 3085 df-dif 3159 df-un 3161 df-in 3163 df-ss 3170 df-nul 3451 df-if 3562 df-pw 3607 df-sn 3628 df-pr 3629 df-tp 3630 df-op 3631 df-uni 3840 df-int 3875 df-iun 3918 df-disj 4011 df-br 4034 df-opab 4095 df-mpt 4096 df-tr 4132 df-id 4328 df-po 4331 df-iso 4332 df-iord 4401 df-on 4403 df-ilim 4404 df-suc 4406 df-iom 4627 df-xp 4669 df-rel 4670 df-cnv 4671 df-co 4672 df-dm 4673 df-rn 4674 df-res 4675 df-ima 4676 df-iota 5219 df-fun 5260 df-fn 5261 df-f 5262 df-f1 5263 df-fo 5264 df-f1o 5265 df-fv 5266 df-isom 5267 df-riota 5877 df-ov 5925 df-oprab 5926 df-mpo 5927 df-of 6135 df-1st 6198 df-2nd 6199 df-tpos 6303 df-recs 6363 df-irdg 6428 df-frec 6449 df-1o 6474 df-2o 6475 df-oadd 6478 df-er 6592 df-ec 6594 df-qs 6598 df-map 6709 df-en 6800 df-dom 6801 df-fin 6802 df-sup 7048 df-inf 7049 df-pnf 8061 df-mnf 8062 df-xr 8063 df-ltxr 8064 df-le 8065 df-sub 8197 df-neg 8198 df-reap 8599 df-ap 8606 df-div 8697 df-inn 8988 df-2 9046 df-3 9047 df-4 9048 df-5 9049 df-6 9050 df-7 9051 df-8 9052 df-9 9053 df-n0 9247 df-z 9324 df-dec 9455 df-uz 9599 df-q 9691 df-rp 9726 df-fz 10081 df-fzo 10215 df-fl 10345 df-mod 10400 df-seqfrec 10525 df-exp 10616 df-ihash 10853 df-cj 10992 df-re 10993 df-im 10994 df-rsqrt 11148 df-abs 11149 df-clim 11428 df-sumdc 11503 df-proddc 11700 df-dvds 11937 df-gcd 12086 df-prm 12252 df-phi 12355 df-pc 12430 df-struct 12656 df-ndx 12657 df-slot 12658 df-base 12660 df-sets 12661 df-iress 12662 df-plusg 12744 df-mulr 12745 df-starv 12746 df-sca 12747 df-vsca 12748 df-ip 12749 df-tset 12750 df-ple 12751 df-ds 12753 df-unif 12754 df-0g 12905 df-igsum 12906 df-topgen 12907 df-iimas 12921 df-qus 12922 df-mgm 12975 df-sgrp 13021 df-mnd 13034 df-mhm 13067 df-submnd 13068 df-grp 13111 df-minusg 13112 df-sbg 13113 df-mulg 13226 df-subg 13276 df-nsg 13277 df-eqg 13278 df-ghm 13347 df-cmn 13392 df-abl 13393 df-mgp 13453 df-rng 13465 df-ur 13492 df-srg 13496 df-ring 13530 df-cring 13531 df-oppr 13600 df-dvdsr 13621 df-unit 13622 df-invr 13653 df-dvr 13664 df-rhm 13684 df-nzr 13712 df-subrg 13751 df-domn 13791 df-idom 13792 df-lmod 13821 df-lssm 13885 df-lsp 13919 df-sra 13967 df-rgmod 13968 df-lidl 14001 df-rsp 14002 df-2idl 14032 df-bl 14078 df-mopn 14079 df-fg 14081 df-metu 14082 df-cnfld 14089 df-zring 14123 df-zrh 14146 df-zn 14148 df-lgs 15206 |
This theorem is referenced by: lgsquad2 15291 |
Copyright terms: Public domain | W3C validator |