MPE Home Metamath Proof Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  MPE Home  >  Th. List  >  eucalgval Structured version   Visualization version   GIF version

Theorem eucalgval 15295
Description: Euclid's Algorithm eucalg 15300 computes the greatest common divisor of two nonnegative integers by repeatedly replacing the larger of them with its remainder modulo the smaller until the remainder is 0.

The value of the step function 𝐸 for Euclid's Algorithm. (Contributed by Paul Chapman, 31-Mar-2011.) (Revised by Mario Carneiro, 28-May-2014.)

Hypothesis
Ref Expression
eucalgval.1 𝐸 = (𝑥 ∈ ℕ0, 𝑦 ∈ ℕ0 ↦ if(𝑦 = 0, ⟨𝑥, 𝑦⟩, ⟨𝑦, (𝑥 mod 𝑦)⟩))
Assertion
Ref Expression
eucalgval (𝑋 ∈ (ℕ0 × ℕ0) → (𝐸𝑋) = if((2nd𝑋) = 0, 𝑋, ⟨(2nd𝑋), ( mod ‘𝑋)⟩))
Distinct variable group:   𝑥,𝑦,𝑋
Allowed substitution hints:   𝐸(𝑥,𝑦)

Proof of Theorem eucalgval
StepHypRef Expression
1 df-ov 6653 . . 3 ((1st𝑋)𝐸(2nd𝑋)) = (𝐸‘⟨(1st𝑋), (2nd𝑋)⟩)
2 xp1st 7198 . . . 4 (𝑋 ∈ (ℕ0 × ℕ0) → (1st𝑋) ∈ ℕ0)
3 xp2nd 7199 . . . 4 (𝑋 ∈ (ℕ0 × ℕ0) → (2nd𝑋) ∈ ℕ0)
4 eucalgval.1 . . . . 5 𝐸 = (𝑥 ∈ ℕ0, 𝑦 ∈ ℕ0 ↦ if(𝑦 = 0, ⟨𝑥, 𝑦⟩, ⟨𝑦, (𝑥 mod 𝑦)⟩))
54eucalgval2 15294 . . . 4 (((1st𝑋) ∈ ℕ0 ∧ (2nd𝑋) ∈ ℕ0) → ((1st𝑋)𝐸(2nd𝑋)) = if((2nd𝑋) = 0, ⟨(1st𝑋), (2nd𝑋)⟩, ⟨(2nd𝑋), ((1st𝑋) mod (2nd𝑋))⟩))
62, 3, 5syl2anc 693 . . 3 (𝑋 ∈ (ℕ0 × ℕ0) → ((1st𝑋)𝐸(2nd𝑋)) = if((2nd𝑋) = 0, ⟨(1st𝑋), (2nd𝑋)⟩, ⟨(2nd𝑋), ((1st𝑋) mod (2nd𝑋))⟩))
71, 6syl5eqr 2670 . 2 (𝑋 ∈ (ℕ0 × ℕ0) → (𝐸‘⟨(1st𝑋), (2nd𝑋)⟩) = if((2nd𝑋) = 0, ⟨(1st𝑋), (2nd𝑋)⟩, ⟨(2nd𝑋), ((1st𝑋) mod (2nd𝑋))⟩))
8 1st2nd2 7205 . . 3 (𝑋 ∈ (ℕ0 × ℕ0) → 𝑋 = ⟨(1st𝑋), (2nd𝑋)⟩)
98fveq2d 6195 . 2 (𝑋 ∈ (ℕ0 × ℕ0) → (𝐸𝑋) = (𝐸‘⟨(1st𝑋), (2nd𝑋)⟩))
108fveq2d 6195 . . . . 5 (𝑋 ∈ (ℕ0 × ℕ0) → ( mod ‘𝑋) = ( mod ‘⟨(1st𝑋), (2nd𝑋)⟩))
11 df-ov 6653 . . . . 5 ((1st𝑋) mod (2nd𝑋)) = ( mod ‘⟨(1st𝑋), (2nd𝑋)⟩)
1210, 11syl6eqr 2674 . . . 4 (𝑋 ∈ (ℕ0 × ℕ0) → ( mod ‘𝑋) = ((1st𝑋) mod (2nd𝑋)))
1312opeq2d 4409 . . 3 (𝑋 ∈ (ℕ0 × ℕ0) → ⟨(2nd𝑋), ( mod ‘𝑋)⟩ = ⟨(2nd𝑋), ((1st𝑋) mod (2nd𝑋))⟩)
148, 13ifeq12d 4106 . 2 (𝑋 ∈ (ℕ0 × ℕ0) → if((2nd𝑋) = 0, 𝑋, ⟨(2nd𝑋), ( mod ‘𝑋)⟩) = if((2nd𝑋) = 0, ⟨(1st𝑋), (2nd𝑋)⟩, ⟨(2nd𝑋), ((1st𝑋) mod (2nd𝑋))⟩))
157, 9, 143eqtr4d 2666 1 (𝑋 ∈ (ℕ0 × ℕ0) → (𝐸𝑋) = if((2nd𝑋) = 0, 𝑋, ⟨(2nd𝑋), ( mod ‘𝑋)⟩))
Colors of variables: wff setvar class
Syntax hints:  wi 4   = wceq 1483  wcel 1990  ifcif 4086  cop 4183   × cxp 5112  cfv 5888  (class class class)co 6650  cmpt2 6652  1st c1st 7166  2nd c2nd 7167  0cc0 9936  0cn0 11292   mod cmo 12668
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1722  ax-4 1737  ax-5 1839  ax-6 1888  ax-7 1935  ax-8 1992  ax-9 1999  ax-10 2019  ax-11 2034  ax-12 2047  ax-13 2246  ax-ext 2602  ax-sep 4781  ax-nul 4789  ax-pow 4843  ax-pr 4906  ax-un 6949
This theorem depends on definitions:  df-bi 197  df-or 385  df-an 386  df-3an 1039  df-tru 1486  df-ex 1705  df-nf 1710  df-sb 1881  df-eu 2474  df-mo 2475  df-clab 2609  df-cleq 2615  df-clel 2618  df-nfc 2753  df-ral 2917  df-rex 2918  df-rab 2921  df-v 3202  df-sbc 3436  df-dif 3577  df-un 3579  df-in 3581  df-ss 3588  df-nul 3916  df-if 4087  df-sn 4178  df-pr 4180  df-op 4184  df-uni 4437  df-br 4654  df-opab 4713  df-mpt 4730  df-id 5024  df-xp 5120  df-rel 5121  df-cnv 5122  df-co 5123  df-dm 5124  df-rn 5125  df-iota 5851  df-fun 5890  df-fv 5896  df-ov 6653  df-oprab 6654  df-mpt2 6655  df-1st 7168  df-2nd 7169
This theorem is referenced by:  eucalginv  15297  eucalglt  15298
  Copyright terms: Public domain W3C validator