ILE Home Intuitionistic Logic Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  ILE Home  >  Th. List  >  bezout GIF version

Theorem bezout 11427
Description: Bézout's identity: For any integers 𝐴 and 𝐵, there are integers 𝑥, 𝑦 such that (𝐴 gcd 𝐵) = 𝐴 · 𝑥 + 𝐵 · 𝑦. This is Metamath 100 proof #60.

The proof is constructive, in the sense that it applies the Extended Euclidian Algorithm to constuct a number which can be shown to be (𝐴 gcd 𝐵) and which satisfies the rest of the theorem. In the presence of excluded middle, it is common to prove Bézout's identity by taking the smallest number which satisfies the Bézout condition, and showing it is the greatest common divisor. But we do not have the ability to show that number exists other than by providing a way to determine it. (Contributed by Mario Carneiro, 22-Feb-2014.)

Assertion
Ref Expression
bezout ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
Distinct variable groups:   𝑥,𝐴,𝑦   𝑥,𝐵,𝑦

Proof of Theorem bezout
Dummy variables 𝑑 𝑤 𝑧 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 bezoutlembi 11421 . 2 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ∃𝑑 ∈ ℕ0 (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
2 simprrr 508 . . 3 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
3 nfv 1473 . . . . 5 𝑥(𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ)
4 nfv 1473 . . . . . 6 𝑥 𝑑 ∈ ℕ0
5 nfv 1473 . . . . . . 7 𝑥𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))
6 nfre1 2430 . . . . . . 7 𝑥𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
75, 6nfan 1509 . . . . . 6 𝑥(∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
84, 7nfan 1509 . . . . 5 𝑥(𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
93, 8nfan 1509 . . . 4 𝑥((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))))
10 nfv 1473 . . . . . 6 𝑦(𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ)
11 nfv 1473 . . . . . . 7 𝑦 𝑑 ∈ ℕ0
12 nfv 1473 . . . . . . . 8 𝑦𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))
13 nfcv 2235 . . . . . . . . 9 𝑦
14 nfre1 2430 . . . . . . . . 9 𝑦𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
1513, 14nfrexya 2428 . . . . . . . 8 𝑦𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
1612, 15nfan 1509 . . . . . . 7 𝑦(∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
1711, 16nfan 1509 . . . . . 6 𝑦(𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
1810, 17nfan 1509 . . . . 5 𝑦((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))))
19 dfgcd3 11426 . . . . . . . 8 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝐴 gcd 𝐵) = (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))))
2019adantr 271 . . . . . . 7 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝐴 gcd 𝐵) = (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))))
21 simprrl 507 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)))
22 simprl 499 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → 𝑑 ∈ ℕ0)
23 simpll 497 . . . . . . . . . 10 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → 𝐴 ∈ ℤ)
24 simplr 498 . . . . . . . . . 10 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → 𝐵 ∈ ℤ)
2523, 24, 22, 21bezoutlemeu 11423 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃!𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)))
26 breq2 3871 . . . . . . . . . . . 12 (𝑤 = 𝑑 → (𝑧𝑤𝑧𝑑))
2726bibi1d 232 . . . . . . . . . . 11 (𝑤 = 𝑑 → ((𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))))
2827ralbidv 2391 . . . . . . . . . 10 (𝑤 = 𝑑 → (∀𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)) ↔ ∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))))
2928riota2 5668 . . . . . . . . 9 ((𝑑 ∈ ℕ0 ∧ ∃!𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) → (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑))
3022, 25, 29syl2anc 404 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑))
3121, 30mpbid 146 . . . . . . 7 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑)
3220, 31eqtrd 2127 . . . . . 6 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝐴 gcd 𝐵) = 𝑑)
3332eqeq1d 2103 . . . . 5 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ((𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
3418, 33rexbid 2390 . . . 4 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
359, 34rexbid 2390 . . 3 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
362, 35mpbird 166 . 2 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
371, 36rexlimddv 2507 1 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
Colors of variables: wff set class
Syntax hints:  wi 4  wa 103  wb 104   = wceq 1296  wcel 1445  wral 2370  wrex 2371  ∃!wreu 2372   class class class wbr 3867  crio 5645  (class class class)co 5690   + caddc 7450   · cmul 7452  0cn0 8771  cz 8848  cdvds 11223   gcd cgcd 11365
This theorem was proved from axioms:  ax-1 5  ax-2 6  ax-mp 7  ax-ia1 105  ax-ia2 106  ax-ia3 107  ax-in1 582  ax-in2 583  ax-io 668  ax-5 1388  ax-7 1389  ax-gen 1390  ax-ie1 1434  ax-ie2 1435  ax-8 1447  ax-10 1448  ax-11 1449  ax-i12 1450  ax-bndl 1451  ax-4 1452  ax-13 1456  ax-14 1457  ax-17 1471  ax-i9 1475  ax-ial 1479  ax-i5r 1480  ax-ext 2077  ax-coll 3975  ax-sep 3978  ax-nul 3986  ax-pow 4030  ax-pr 4060  ax-un 4284  ax-setind 4381  ax-iinf 4431  ax-cnex 7533  ax-resscn 7534  ax-1cn 7535  ax-1re 7536  ax-icn 7537  ax-addcl 7538  ax-addrcl 7539  ax-mulcl 7540  ax-mulrcl 7541  ax-addcom 7542  ax-mulcom 7543  ax-addass 7544  ax-mulass 7545  ax-distr 7546  ax-i2m1 7547  ax-0lt1 7548  ax-1rid 7549  ax-0id 7550  ax-rnegex 7551  ax-precex 7552  ax-cnre 7553  ax-pre-ltirr 7554  ax-pre-ltwlin 7555  ax-pre-lttrn 7556  ax-pre-apti 7557  ax-pre-ltadd 7558  ax-pre-mulgt0 7559  ax-pre-mulext 7560  ax-arch 7561  ax-caucvg 7562
This theorem depends on definitions:  df-bi 116  df-dc 784  df-3or 928  df-3an 929  df-tru 1299  df-fal 1302  df-nf 1402  df-sb 1700  df-eu 1958  df-mo 1959  df-clab 2082  df-cleq 2088  df-clel 2091  df-nfc 2224  df-ne 2263  df-nel 2358  df-ral 2375  df-rex 2376  df-reu 2377  df-rmo 2378  df-rab 2379  df-v 2635  df-sbc 2855  df-csb 2948  df-dif 3015  df-un 3017  df-in 3019  df-ss 3026  df-nul 3303  df-if 3414  df-pw 3451  df-sn 3472  df-pr 3473  df-op 3475  df-uni 3676  df-int 3711  df-iun 3754  df-br 3868  df-opab 3922  df-mpt 3923  df-tr 3959  df-id 4144  df-po 4147  df-iso 4148  df-iord 4217  df-on 4219  df-ilim 4220  df-suc 4222  df-iom 4434  df-xp 4473  df-rel 4474  df-cnv 4475  df-co 4476  df-dm 4477  df-rn 4478  df-res 4479  df-ima 4480  df-iota 5014  df-fun 5051  df-fn 5052  df-f 5053  df-f1 5054  df-fo 5055  df-f1o 5056  df-fv 5057  df-riota 5646  df-ov 5693  df-oprab 5694  df-mpt2 5695  df-1st 5949  df-2nd 5950  df-recs 6108  df-frec 6194  df-sup 6759  df-pnf 7621  df-mnf 7622  df-xr 7623  df-ltxr 7624  df-le 7625  df-sub 7752  df-neg 7753  df-reap 8149  df-ap 8156  df-div 8237  df-inn 8521  df-2 8579  df-3 8580  df-4 8581  df-n0 8772  df-z 8849  df-uz 9119  df-q 9204  df-rp 9234  df-fz 9574  df-fzo 9703  df-fl 9826  df-mod 9879  df-seqfrec 10001  df-exp 10070  df-cj 10391  df-re 10392  df-im 10393  df-rsqrt 10546  df-abs 10547  df-dvds 11224  df-gcd 11366
This theorem is referenced by:  dvdsgcd  11428  dvdsmulgcd  11441  lcmgcdlem  11486  divgcdcoprm0  11510
  Copyright terms: Public domain W3C validator