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

Theorem bezout 11944
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 11938 . 2 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ∃𝑑 ∈ ℕ0 (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
2 simprrr 530 . . 3 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
3 nfv 1516 . . . . 5 𝑥(𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ)
4 nfv 1516 . . . . . 6 𝑥 𝑑 ∈ ℕ0
5 nfv 1516 . . . . . . 7 𝑥𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))
6 nfre1 2509 . . . . . . 7 𝑥𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
75, 6nfan 1553 . . . . . 6 𝑥(∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
84, 7nfan 1553 . . . . 5 𝑥(𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
93, 8nfan 1553 . . . 4 𝑥((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))))
10 nfv 1516 . . . . . 6 𝑦(𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ)
11 nfv 1516 . . . . . . 7 𝑦 𝑑 ∈ ℕ0
12 nfv 1516 . . . . . . . 8 𝑦𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))
13 nfcv 2308 . . . . . . . . 9 𝑦
14 nfre1 2509 . . . . . . . . 9 𝑦𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
1513, 14nfrexya 2507 . . . . . . . 8 𝑦𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
1612, 15nfan 1553 . . . . . . 7 𝑦(∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
1711, 16nfan 1553 . . . . . 6 𝑦(𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
1810, 17nfan 1553 . . . . 5 𝑦((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))))
19 dfgcd3 11943 . . . . . . . 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 11940 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃!𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)))
26 breq2 3986 . . . . . . . . . . . 12 (𝑤 = 𝑑 → (𝑧𝑤𝑧𝑑))
2726bibi1d 232 . . . . . . . . . . 11 (𝑤 = 𝑑 → ((𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))))
2827ralbidv 2466 . . . . . . . . . 10 (𝑤 = 𝑑 → (∀𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)) ↔ ∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))))
2928riota2 5820 . . . . . . . . 9 ((𝑑 ∈ ℕ0 ∧ ∃!𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) → (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑))
3022, 25, 29syl2anc 409 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑))
3121, 30mpbid 146 . . . . . . 7 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑)
3220, 31eqtrd 2198 . . . . . 6 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝐴 gcd 𝐵) = 𝑑)
3332eqeq1d 2174 . . . . 5 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ((𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
3418, 33rexbid 2465 . . . 4 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
359, 34rexbid 2465 . . 3 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
362, 35mpbird 166 . 2 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
371, 36rexlimddv 2588 1 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
Colors of variables: wff set class
Syntax hints:  wi 4  wa 103  wb 104   = wceq 1343  wcel 2136  wral 2444  wrex 2445  ∃!wreu 2446   class class class wbr 3982  crio 5797  (class class class)co 5842   + caddc 7756   · cmul 7758  0cn0 9114  cz 9191  cdvds 11727   gcd cgcd 11875
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 1435  ax-7 1436  ax-gen 1437  ax-ie1 1481  ax-ie2 1482  ax-8 1492  ax-10 1493  ax-11 1494  ax-i12 1495  ax-bndl 1497  ax-4 1498  ax-17 1514  ax-i9 1518  ax-ial 1522  ax-i5r 1523  ax-13 2138  ax-14 2139  ax-ext 2147  ax-coll 4097  ax-sep 4100  ax-nul 4108  ax-pow 4153  ax-pr 4187  ax-un 4411  ax-setind 4514  ax-iinf 4565  ax-cnex 7844  ax-resscn 7845  ax-1cn 7846  ax-1re 7847  ax-icn 7848  ax-addcl 7849  ax-addrcl 7850  ax-mulcl 7851  ax-mulrcl 7852  ax-addcom 7853  ax-mulcom 7854  ax-addass 7855  ax-mulass 7856  ax-distr 7857  ax-i2m1 7858  ax-0lt1 7859  ax-1rid 7860  ax-0id 7861  ax-rnegex 7862  ax-precex 7863  ax-cnre 7864  ax-pre-ltirr 7865  ax-pre-ltwlin 7866  ax-pre-lttrn 7867  ax-pre-apti 7868  ax-pre-ltadd 7869  ax-pre-mulgt0 7870  ax-pre-mulext 7871  ax-arch 7872  ax-caucvg 7873
This theorem depends on definitions:  df-bi 116  df-dc 825  df-3or 969  df-3an 970  df-tru 1346  df-fal 1349  df-nf 1449  df-sb 1751  df-eu 2017  df-mo 2018  df-clab 2152  df-cleq 2158  df-clel 2161  df-nfc 2297  df-ne 2337  df-nel 2432  df-ral 2449  df-rex 2450  df-reu 2451  df-rmo 2452  df-rab 2453  df-v 2728  df-sbc 2952  df-csb 3046  df-dif 3118  df-un 3120  df-in 3122  df-ss 3129  df-nul 3410  df-if 3521  df-pw 3561  df-sn 3582  df-pr 3583  df-op 3585  df-uni 3790  df-int 3825  df-iun 3868  df-br 3983  df-opab 4044  df-mpt 4045  df-tr 4081  df-id 4271  df-po 4274  df-iso 4275  df-iord 4344  df-on 4346  df-ilim 4347  df-suc 4349  df-iom 4568  df-xp 4610  df-rel 4611  df-cnv 4612  df-co 4613  df-dm 4614  df-rn 4615  df-res 4616  df-ima 4617  df-iota 5153  df-fun 5190  df-fn 5191  df-f 5192  df-f1 5193  df-fo 5194  df-f1o 5195  df-fv 5196  df-riota 5798  df-ov 5845  df-oprab 5846  df-mpo 5847  df-1st 6108  df-2nd 6109  df-recs 6273  df-frec 6359  df-sup 6949  df-pnf 7935  df-mnf 7936  df-xr 7937  df-ltxr 7938  df-le 7939  df-sub 8071  df-neg 8072  df-reap 8473  df-ap 8480  df-div 8569  df-inn 8858  df-2 8916  df-3 8917  df-4 8918  df-n0 9115  df-z 9192  df-uz 9467  df-q 9558  df-rp 9590  df-fz 9945  df-fzo 10078  df-fl 10205  df-mod 10258  df-seqfrec 10381  df-exp 10455  df-cj 10784  df-re 10785  df-im 10786  df-rsqrt 10940  df-abs 10941  df-dvds 11728  df-gcd 11876
This theorem is referenced by:  dvdsgcd  11945  dvdsmulgcd  11958  lcmgcdlem  12009  divgcdcoprm0  12033
  Copyright terms: Public domain W3C validator