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

Theorem bezout 11688
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 11682 . 2 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ∃𝑑 ∈ ℕ0 (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
2 simprrr 529 . . 3 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
3 nfv 1508 . . . . 5 𝑥(𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ)
4 nfv 1508 . . . . . 6 𝑥 𝑑 ∈ ℕ0
5 nfv 1508 . . . . . . 7 𝑥𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))
6 nfre1 2474 . . . . . . 7 𝑥𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
75, 6nfan 1544 . . . . . 6 𝑥(∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
84, 7nfan 1544 . . . . 5 𝑥(𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
93, 8nfan 1544 . . . 4 𝑥((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))))
10 nfv 1508 . . . . . 6 𝑦(𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ)
11 nfv 1508 . . . . . . 7 𝑦 𝑑 ∈ ℕ0
12 nfv 1508 . . . . . . . 8 𝑦𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))
13 nfcv 2279 . . . . . . . . 9 𝑦
14 nfre1 2474 . . . . . . . . 9 𝑦𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
1513, 14nfrexya 2472 . . . . . . . 8 𝑦𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
1612, 15nfan 1544 . . . . . . 7 𝑦(∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
1711, 16nfan 1544 . . . . . 6 𝑦(𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
1810, 17nfan 1544 . . . . 5 𝑦((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))))
19 dfgcd3 11687 . . . . . . . 8 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝐴 gcd 𝐵) = (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))))
2019adantr 274 . . . . . . 7 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝐴 gcd 𝐵) = (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))))
21 simprrl 528 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)))
22 simprl 520 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → 𝑑 ∈ ℕ0)
23 simpll 518 . . . . . . . . . 10 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → 𝐴 ∈ ℤ)
24 simplr 519 . . . . . . . . . 10 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → 𝐵 ∈ ℤ)
2523, 24, 22, 21bezoutlemeu 11684 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃!𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)))
26 breq2 3928 . . . . . . . . . . . 12 (𝑤 = 𝑑 → (𝑧𝑤𝑧𝑑))
2726bibi1d 232 . . . . . . . . . . 11 (𝑤 = 𝑑 → ((𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))))
2827ralbidv 2435 . . . . . . . . . 10 (𝑤 = 𝑑 → (∀𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)) ↔ ∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))))
2928riota2 5745 . . . . . . . . 9 ((𝑑 ∈ ℕ0 ∧ ∃!𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) → (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑))
3022, 25, 29syl2anc 408 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑))
3121, 30mpbid 146 . . . . . . 7 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑)
3220, 31eqtrd 2170 . . . . . 6 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝐴 gcd 𝐵) = 𝑑)
3332eqeq1d 2146 . . . . 5 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ((𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
3418, 33rexbid 2434 . . . 4 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
359, 34rexbid 2434 . . 3 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
362, 35mpbird 166 . 2 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
371, 36rexlimddv 2552 1 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
Colors of variables: wff set class
Syntax hints:  wi 4  wa 103  wb 104   = wceq 1331  wcel 1480  wral 2414  wrex 2415  ∃!wreu 2416   class class class wbr 3924  crio 5722  (class class class)co 5767   + caddc 7616   · cmul 7618  0cn0 8970  cz 9047  cdvds 11482   gcd cgcd 11624
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-ia1 105  ax-ia2 106  ax-ia3 107  ax-in1 603  ax-in2 604  ax-io 698  ax-5 1423  ax-7 1424  ax-gen 1425  ax-ie1 1469  ax-ie2 1470  ax-8 1482  ax-10 1483  ax-11 1484  ax-i12 1485  ax-bndl 1486  ax-4 1487  ax-13 1491  ax-14 1492  ax-17 1506  ax-i9 1510  ax-ial 1514  ax-i5r 1515  ax-ext 2119  ax-coll 4038  ax-sep 4041  ax-nul 4049  ax-pow 4093  ax-pr 4126  ax-un 4350  ax-setind 4447  ax-iinf 4497  ax-cnex 7704  ax-resscn 7705  ax-1cn 7706  ax-1re 7707  ax-icn 7708  ax-addcl 7709  ax-addrcl 7710  ax-mulcl 7711  ax-mulrcl 7712  ax-addcom 7713  ax-mulcom 7714  ax-addass 7715  ax-mulass 7716  ax-distr 7717  ax-i2m1 7718  ax-0lt1 7719  ax-1rid 7720  ax-0id 7721  ax-rnegex 7722  ax-precex 7723  ax-cnre 7724  ax-pre-ltirr 7725  ax-pre-ltwlin 7726  ax-pre-lttrn 7727  ax-pre-apti 7728  ax-pre-ltadd 7729  ax-pre-mulgt0 7730  ax-pre-mulext 7731  ax-arch 7732  ax-caucvg 7733
This theorem depends on definitions:  df-bi 116  df-dc 820  df-3or 963  df-3an 964  df-tru 1334  df-fal 1337  df-nf 1437  df-sb 1736  df-eu 2000  df-mo 2001  df-clab 2124  df-cleq 2130  df-clel 2133  df-nfc 2268  df-ne 2307  df-nel 2402  df-ral 2419  df-rex 2420  df-reu 2421  df-rmo 2422  df-rab 2423  df-v 2683  df-sbc 2905  df-csb 2999  df-dif 3068  df-un 3070  df-in 3072  df-ss 3079  df-nul 3359  df-if 3470  df-pw 3507  df-sn 3528  df-pr 3529  df-op 3531  df-uni 3732  df-int 3767  df-iun 3810  df-br 3925  df-opab 3985  df-mpt 3986  df-tr 4022  df-id 4210  df-po 4213  df-iso 4214  df-iord 4283  df-on 4285  df-ilim 4286  df-suc 4288  df-iom 4500  df-xp 4540  df-rel 4541  df-cnv 4542  df-co 4543  df-dm 4544  df-rn 4545  df-res 4546  df-ima 4547  df-iota 5083  df-fun 5120  df-fn 5121  df-f 5122  df-f1 5123  df-fo 5124  df-f1o 5125  df-fv 5126  df-riota 5723  df-ov 5770  df-oprab 5771  df-mpo 5772  df-1st 6031  df-2nd 6032  df-recs 6195  df-frec 6281  df-sup 6864  df-pnf 7795  df-mnf 7796  df-xr 7797  df-ltxr 7798  df-le 7799  df-sub 7928  df-neg 7929  df-reap 8330  df-ap 8337  df-div 8426  df-inn 8714  df-2 8772  df-3 8773  df-4 8774  df-n0 8971  df-z 9048  df-uz 9320  df-q 9405  df-rp 9435  df-fz 9784  df-fzo 9913  df-fl 10036  df-mod 10089  df-seqfrec 10212  df-exp 10286  df-cj 10607  df-re 10608  df-im 10609  df-rsqrt 10763  df-abs 10764  df-dvds 11483  df-gcd 11625
This theorem is referenced by:  dvdsgcd  11689  dvdsmulgcd  11702  lcmgcdlem  11747  divgcdcoprm0  11771
  Copyright terms: Public domain W3C validator