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

Theorem eucalgval 15133
Description: Euclid's Algorithm eucalg 15138 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 6552 . . 3 ((1st𝑋)𝐸(2nd𝑋)) = (𝐸‘⟨(1st𝑋), (2nd𝑋)⟩)
2 xp1st 7089 . . . 4 (𝑋 ∈ (ℕ0 × ℕ0) → (1st𝑋) ∈ ℕ0)
3 xp2nd 7090 . . . 4 (𝑋 ∈ (ℕ0 × ℕ0) → (2nd𝑋) ∈ ℕ0)
4 eucalgval.1 . . . . 5 𝐸 = (𝑥 ∈ ℕ0, 𝑦 ∈ ℕ0 ↦ if(𝑦 = 0, ⟨𝑥, 𝑦⟩, ⟨𝑦, (𝑥 mod 𝑦)⟩))
54eucalgval2 15132 . . . 4 (((1st𝑋) ∈ ℕ0 ∧ (2nd𝑋) ∈ ℕ0) → ((1st𝑋)𝐸(2nd𝑋)) = if((2nd𝑋) = 0, ⟨(1st𝑋), (2nd𝑋)⟩, ⟨(2nd𝑋), ((1st𝑋) mod (2nd𝑋))⟩))
62, 3, 5syl2anc 691 . . 3 (𝑋 ∈ (ℕ0 × ℕ0) → ((1st𝑋)𝐸(2nd𝑋)) = if((2nd𝑋) = 0, ⟨(1st𝑋), (2nd𝑋)⟩, ⟨(2nd𝑋), ((1st𝑋) mod (2nd𝑋))⟩))
71, 6syl5eqr 2658 . 2 (𝑋 ∈ (ℕ0 × ℕ0) → (𝐸‘⟨(1st𝑋), (2nd𝑋)⟩) = if((2nd𝑋) = 0, ⟨(1st𝑋), (2nd𝑋)⟩, ⟨(2nd𝑋), ((1st𝑋) mod (2nd𝑋))⟩))
8 1st2nd2 7096 . . 3 (𝑋 ∈ (ℕ0 × ℕ0) → 𝑋 = ⟨(1st𝑋), (2nd𝑋)⟩)
98fveq2d 6107 . 2 (𝑋 ∈ (ℕ0 × ℕ0) → (𝐸𝑋) = (𝐸‘⟨(1st𝑋), (2nd𝑋)⟩))
108fveq2d 6107 . . . . 5 (𝑋 ∈ (ℕ0 × ℕ0) → ( mod ‘𝑋) = ( mod ‘⟨(1st𝑋), (2nd𝑋)⟩))
11 df-ov 6552 . . . . 5 ((1st𝑋) mod (2nd𝑋)) = ( mod ‘⟨(1st𝑋), (2nd𝑋)⟩)
1210, 11syl6eqr 2662 . . . 4 (𝑋 ∈ (ℕ0 × ℕ0) → ( mod ‘𝑋) = ((1st𝑋) mod (2nd𝑋)))
1312opeq2d 4347 . . 3 (𝑋 ∈ (ℕ0 × ℕ0) → ⟨(2nd𝑋), ( mod ‘𝑋)⟩ = ⟨(2nd𝑋), ((1st𝑋) mod (2nd𝑋))⟩)
148, 13ifeq12d 4056 . 2 (𝑋 ∈ (ℕ0 × ℕ0) → if((2nd𝑋) = 0, 𝑋, ⟨(2nd𝑋), ( mod ‘𝑋)⟩) = if((2nd𝑋) = 0, ⟨(1st𝑋), (2nd𝑋)⟩, ⟨(2nd𝑋), ((1st𝑋) mod (2nd𝑋))⟩))
157, 9, 143eqtr4d 2654 1 (𝑋 ∈ (ℕ0 × ℕ0) → (𝐸𝑋) = if((2nd𝑋) = 0, 𝑋, ⟨(2nd𝑋), ( mod ‘𝑋)⟩))
Colors of variables: wff setvar class
Syntax hints:  wi 4   = wceq 1475  wcel 1977  ifcif 4036  cop 4131   × cxp 5036  cfv 5804  (class class class)co 6549  cmpt2 6551  1st c1st 7057  2nd c2nd 7058  0cc0 9815  0cn0 11169   mod cmo 12530
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1713  ax-4 1728  ax-5 1827  ax-6 1875  ax-7 1922  ax-8 1979  ax-9 1986  ax-10 2006  ax-11 2021  ax-12 2034  ax-13 2234  ax-ext 2590  ax-sep 4709  ax-nul 4717  ax-pow 4769  ax-pr 4833  ax-un 6847
This theorem depends on definitions:  df-bi 196  df-or 384  df-an 385  df-3an 1033  df-tru 1478  df-ex 1696  df-nf 1701  df-sb 1868  df-eu 2462  df-mo 2463  df-clab 2597  df-cleq 2603  df-clel 2606  df-nfc 2740  df-ral 2901  df-rex 2902  df-rab 2905  df-v 3175  df-sbc 3403  df-dif 3543  df-un 3545  df-in 3547  df-ss 3554  df-nul 3875  df-if 4037  df-sn 4126  df-pr 4128  df-op 4132  df-uni 4373  df-br 4584  df-opab 4644  df-mpt 4645  df-id 4953  df-xp 5044  df-rel 5045  df-cnv 5046  df-co 5047  df-dm 5048  df-rn 5049  df-iota 5768  df-fun 5806  df-fv 5812  df-ov 6552  df-oprab 6553  df-mpt2 6554  df-1st 7059  df-2nd 7060
This theorem is referenced by:  eucalginv  15135  eucalglt  15136
  Copyright terms: Public domain W3C validator