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

Theorem bezout 11966
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 11960 . 2 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ∃𝑑 ∈ ℕ0 (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
2 simprrr 535 . . 3 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
3 nfv 1521 . . . . 5 𝑥(𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ)
4 nfv 1521 . . . . . 6 𝑥 𝑑 ∈ ℕ0
5 nfv 1521 . . . . . . 7 𝑥𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))
6 nfre1 2513 . . . . . . 7 𝑥𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
75, 6nfan 1558 . . . . . 6 𝑥(∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
84, 7nfan 1558 . . . . 5 𝑥(𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
93, 8nfan 1558 . . . 4 𝑥((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))))
10 nfv 1521 . . . . . 6 𝑦(𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ)
11 nfv 1521 . . . . . . 7 𝑦 𝑑 ∈ ℕ0
12 nfv 1521 . . . . . . . 8 𝑦𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))
13 nfcv 2312 . . . . . . . . 9 𝑦
14 nfre1 2513 . . . . . . . . 9 𝑦𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
1513, 14nfrexya 2511 . . . . . . . 8 𝑦𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))
1612, 15nfan 1558 . . . . . . 7 𝑦(∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
1711, 16nfan 1558 . . . . . 6 𝑦(𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
1810, 17nfan 1558 . . . . 5 𝑦((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))))
19 dfgcd3 11965 . . . . . . . 8 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝐴 gcd 𝐵) = (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))))
2019adantr 274 . . . . . . 7 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝐴 gcd 𝐵) = (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))))
21 simprrl 534 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)))
22 simprl 526 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → 𝑑 ∈ ℕ0)
23 simpll 524 . . . . . . . . . 10 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → 𝐴 ∈ ℤ)
24 simplr 525 . . . . . . . . . 10 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → 𝐵 ∈ ℤ)
2523, 24, 22, 21bezoutlemeu 11962 . . . . . . . . 9 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃!𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)))
26 breq2 3993 . . . . . . . . . . . 12 (𝑤 = 𝑑 → (𝑧𝑤𝑧𝑑))
2726bibi1d 232 . . . . . . . . . . 11 (𝑤 = 𝑑 → ((𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))))
2827ralbidv 2470 . . . . . . . . . 10 (𝑤 = 𝑑 → (∀𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵)) ↔ ∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵))))
2928riota2 5831 . . . . . . . . 9 ((𝑑 ∈ ℕ0 ∧ ∃!𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) → (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑))
3022, 25, 29syl2anc 409 . . . . . . . 8 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ↔ (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑))
3121, 30mpbid 146 . . . . . . 7 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝑤 ∈ ℕ0𝑧 ∈ ℤ (𝑧𝑤 ↔ (𝑧𝐴𝑧𝐵))) = 𝑑)
3220, 31eqtrd 2203 . . . . . 6 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (𝐴 gcd 𝐵) = 𝑑)
3332eqeq1d 2179 . . . . 5 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ((𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
3418, 33rexbid 2469 . . . 4 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
359, 34rexbid 2469 . . 3 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → (∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)) ↔ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))
362, 35mpbird 166 . 2 (((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) ∧ (𝑑 ∈ ℕ0 ∧ (∀𝑧 ∈ ℤ (𝑧𝑑 ↔ (𝑧𝐴𝑧𝐵)) ∧ ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ 𝑑 = ((𝐴 · 𝑥) + (𝐵 · 𝑦))))) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
371, 36rexlimddv 2592 1 ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ∃𝑥 ∈ ℤ ∃𝑦 ∈ ℤ (𝐴 gcd 𝐵) = ((𝐴 · 𝑥) + (𝐵 · 𝑦)))
Colors of variables: wff set class
Syntax hints:  wi 4  wa 103  wb 104   = wceq 1348  wcel 2141  wral 2448  wrex 2449  ∃!wreu 2450   class class class wbr 3989  crio 5808  (class class class)co 5853   + caddc 7777   · cmul 7779  0cn0 9135  cz 9212  cdvds 11749   gcd cgcd 11897
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 609  ax-in2 610  ax-io 704  ax-5 1440  ax-7 1441  ax-gen 1442  ax-ie1 1486  ax-ie2 1487  ax-8 1497  ax-10 1498  ax-11 1499  ax-i12 1500  ax-bndl 1502  ax-4 1503  ax-17 1519  ax-i9 1523  ax-ial 1527  ax-i5r 1528  ax-13 2143  ax-14 2144  ax-ext 2152  ax-coll 4104  ax-sep 4107  ax-nul 4115  ax-pow 4160  ax-pr 4194  ax-un 4418  ax-setind 4521  ax-iinf 4572  ax-cnex 7865  ax-resscn 7866  ax-1cn 7867  ax-1re 7868  ax-icn 7869  ax-addcl 7870  ax-addrcl 7871  ax-mulcl 7872  ax-mulrcl 7873  ax-addcom 7874  ax-mulcom 7875  ax-addass 7876  ax-mulass 7877  ax-distr 7878  ax-i2m1 7879  ax-0lt1 7880  ax-1rid 7881  ax-0id 7882  ax-rnegex 7883  ax-precex 7884  ax-cnre 7885  ax-pre-ltirr 7886  ax-pre-ltwlin 7887  ax-pre-lttrn 7888  ax-pre-apti 7889  ax-pre-ltadd 7890  ax-pre-mulgt0 7891  ax-pre-mulext 7892  ax-arch 7893  ax-caucvg 7894
This theorem depends on definitions:  df-bi 116  df-dc 830  df-3or 974  df-3an 975  df-tru 1351  df-fal 1354  df-nf 1454  df-sb 1756  df-eu 2022  df-mo 2023  df-clab 2157  df-cleq 2163  df-clel 2166  df-nfc 2301  df-ne 2341  df-nel 2436  df-ral 2453  df-rex 2454  df-reu 2455  df-rmo 2456  df-rab 2457  df-v 2732  df-sbc 2956  df-csb 3050  df-dif 3123  df-un 3125  df-in 3127  df-ss 3134  df-nul 3415  df-if 3527  df-pw 3568  df-sn 3589  df-pr 3590  df-op 3592  df-uni 3797  df-int 3832  df-iun 3875  df-br 3990  df-opab 4051  df-mpt 4052  df-tr 4088  df-id 4278  df-po 4281  df-iso 4282  df-iord 4351  df-on 4353  df-ilim 4354  df-suc 4356  df-iom 4575  df-xp 4617  df-rel 4618  df-cnv 4619  df-co 4620  df-dm 4621  df-rn 4622  df-res 4623  df-ima 4624  df-iota 5160  df-fun 5200  df-fn 5201  df-f 5202  df-f1 5203  df-fo 5204  df-f1o 5205  df-fv 5206  df-riota 5809  df-ov 5856  df-oprab 5857  df-mpo 5858  df-1st 6119  df-2nd 6120  df-recs 6284  df-frec 6370  df-sup 6961  df-pnf 7956  df-mnf 7957  df-xr 7958  df-ltxr 7959  df-le 7960  df-sub 8092  df-neg 8093  df-reap 8494  df-ap 8501  df-div 8590  df-inn 8879  df-2 8937  df-3 8938  df-4 8939  df-n0 9136  df-z 9213  df-uz 9488  df-q 9579  df-rp 9611  df-fz 9966  df-fzo 10099  df-fl 10226  df-mod 10279  df-seqfrec 10402  df-exp 10476  df-cj 10806  df-re 10807  df-im 10808  df-rsqrt 10962  df-abs 10963  df-dvds 11750  df-gcd 11898
This theorem is referenced by:  dvdsgcd  11967  dvdsmulgcd  11980  lcmgcdlem  12031  divgcdcoprm0  12055
  Copyright terms: Public domain W3C validator