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

Theorem bezout 11929
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 11923 . 2 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ∃𝑑 ∈ ℕ0 (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
2 simprrr 530 . . 3 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
3 nfv 1515 . . . . 5 𝑥(𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ)
4 nfv 1515 . . . . . 6 𝑥 𝑑 ∈ ℕ0
5 nfv 1515 . . . . . . 7 𝑥𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))
6 nfre1 2507 . . . . . . 7 𝑥𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
75, 6nfan 1552 . . . . . 6 𝑥(∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
84, 7nfan 1552 . . . . 5 𝑥(𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
93, 8nfan 1552 . . . 4 𝑥((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))))
10 nfv 1515 . . . . . 6 𝑦(𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ)
11 nfv 1515 . . . . . . 7 𝑦 𝑑 ∈ ℕ0
12 nfv 1515 . . . . . . . 8 𝑦𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))
13 nfcv 2306 . . . . . . . . 9 𝑦
14 nfre1 2507 . . . . . . . . 9 𝑦𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
1513, 14nfrexya 2505 . . . . . . . 8 𝑦𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
1612, 15nfan 1552 . . . . . . 7 𝑦(∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
1711, 16nfan 1552 . . . . . 6 𝑦(𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
1810, 17nfan 1552 . . . . 5 𝑦((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))))
19 dfgcd3 11928 . . . . . . . 8 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝐴 gcd 𝐵) = (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))))
2019adantr 274 . . . . . . 7 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝐴 gcd 𝐵) = (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))))
21 simprrl 529 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)))
22 simprl 521 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → 𝑑 ∈ ℕ0)
23 simpll 519 . . . . . . . . . 10 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → 𝐴 ∈ ℤ)
24 simplr 520 . . . . . . . . . 10 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → 𝐵 ∈ ℤ)
2523, 24, 22, 21bezoutlemeu 11925 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃!𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)))
26 breq2 3980 . . . . . . . . . . . 12 (𝑤 = 𝑑 → (𝑧𝑤𝑧𝑑))
2726bibi1d 232 . . . . . . . . . . 11 (𝑤 = 𝑑 → ((𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))))
2827ralbidv 2464 . . . . . . . . . 10 (𝑤 = 𝑑 → (∀𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)) ↔ ∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))))
2928riota2 5814 . . . . . . . . 9 ((𝑑 ∈ ℕ0 ∧ ∃!𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) → (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑))
3022, 25, 29syl2anc 409 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑))
3121, 30mpbid 146 . . . . . . 7 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑)
3220, 31eqtrd 2197 . . . . . 6 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝐴 gcd 𝐵) = 𝑑)
3332eqeq1d 2173 . . . . 5 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ((𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
3418, 33rexbid 2463 . . . 4 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
359, 34rexbid 2463 . . 3 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
362, 35mpbird 166 . 2 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
371, 36rexlimddv 2586 1 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
Colors of variables: wff set class
Syntax hints:  wi 4  wa 103  wb 104   = wceq 1342  wcel 2135  wral 2442  wrex 2443  ∃!wreu 2444   class class class wbr 3976  crio 5791  (class class class)co 5836   + caddc 7747   · cmul 7749  0cn0 9105  cz 9182  cdvds 11713   gcd cgcd 11860
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 604  ax-in2 605  ax-io 699  ax-5 1434  ax-7 1435  ax-gen 1436  ax-ie1 1480  ax-ie2 1481  ax-8 1491  ax-10 1492  ax-11 1493  ax-i12 1494  ax-bndl 1496  ax-4 1497  ax-17 1513  ax-i9 1517  ax-ial 1521  ax-i5r 1522  ax-13 2137  ax-14 2138  ax-ext 2146  ax-coll 4091  ax-sep 4094  ax-nul 4102  ax-pow 4147  ax-pr 4181  ax-un 4405  ax-setind 4508  ax-iinf 4559  ax-cnex 7835  ax-resscn 7836  ax-1cn 7837  ax-1re 7838  ax-icn 7839  ax-addcl 7840  ax-addrcl 7841  ax-mulcl 7842  ax-mulrcl 7843  ax-addcom 7844  ax-mulcom 7845  ax-addass 7846  ax-mulass 7847  ax-distr 7848  ax-i2m1 7849  ax-0lt1 7850  ax-1rid 7851  ax-0id 7852  ax-rnegex 7853  ax-precex 7854  ax-cnre 7855  ax-pre-ltirr 7856  ax-pre-ltwlin 7857  ax-pre-lttrn 7858  ax-pre-apti 7859  ax-pre-ltadd 7860  ax-pre-mulgt0 7861  ax-pre-mulext 7862  ax-arch 7863  ax-caucvg 7864
This theorem depends on definitions:  df-bi 116  df-dc 825  df-3or 968  df-3an 969  df-tru 1345  df-fal 1348  df-nf 1448  df-sb 1750  df-eu 2016  df-mo 2017  df-clab 2151  df-cleq 2157  df-clel 2160  df-nfc 2295  df-ne 2335  df-nel 2430  df-ral 2447  df-rex 2448  df-reu 2449  df-rmo 2450  df-rab 2451  df-v 2723  df-sbc 2947  df-csb 3041  df-dif 3113  df-un 3115  df-in 3117  df-ss 3124  df-nul 3405  df-if 3516  df-pw 3555  df-sn 3576  df-pr 3577  df-op 3579  df-uni 3784  df-int 3819  df-iun 3862  df-br 3977  df-opab 4038  df-mpt 4039  df-tr 4075  df-id 4265  df-po 4268  df-iso 4269  df-iord 4338  df-on 4340  df-ilim 4341  df-suc 4343  df-iom 4562  df-xp 4604  df-rel 4605  df-cnv 4606  df-co 4607  df-dm 4608  df-rn 4609  df-res 4610  df-ima 4611  df-iota 5147  df-fun 5184  df-fn 5185  df-f 5186  df-f1 5187  df-fo 5188  df-f1o 5189  df-fv 5190  df-riota 5792  df-ov 5839  df-oprab 5840  df-mpo 5841  df-1st 6100  df-2nd 6101  df-recs 6264  df-frec 6350  df-sup 6940  df-pnf 7926  df-mnf 7927  df-xr 7928  df-ltxr 7929  df-le 7930  df-sub 8062  df-neg 8063  df-reap 8464  df-ap 8471  df-div 8560  df-inn 8849  df-2 8907  df-3 8908  df-4 8909  df-n0 9106  df-z 9183  df-uz 9458  df-q 9549  df-rp 9581  df-fz 9936  df-fzo 10068  df-fl 10195  df-mod 10248  df-seqfrec 10371  df-exp 10445  df-cj 10770  df-re 10771  df-im 10772  df-rsqrt 10926  df-abs 10927  df-dvds 11714  df-gcd 11861
This theorem is referenced by:  dvdsgcd  11930  dvdsmulgcd  11943  lcmgcdlem  11988  divgcdcoprm0  12012
  Copyright terms: Public domain W3C validator