Metamath Proof Explorer


Theorem gcd32

Description: Swap the second and third arguments of a gcd. (Contributed by Scott Fenton, 8-Apr-2014) (Revised by Mario Carneiro, 19-Apr-2014)

Ref Expression
Assertion gcd32 ( ( 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ ∧ 𝐶 ∈ ℤ ) → ( ( 𝐴 gcd 𝐵 ) gcd 𝐶 ) = ( ( 𝐴 gcd 𝐶 ) gcd 𝐵 ) )

Proof

Step Hyp Ref Expression
1 gcdcom ( ( 𝐵 ∈ ℤ ∧ 𝐶 ∈ ℤ ) → ( 𝐵 gcd 𝐶 ) = ( 𝐶 gcd 𝐵 ) )
2 1 3adant1 ( ( 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ ∧ 𝐶 ∈ ℤ ) → ( 𝐵 gcd 𝐶 ) = ( 𝐶 gcd 𝐵 ) )
3 2 oveq2d ( ( 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ ∧ 𝐶 ∈ ℤ ) → ( 𝐴 gcd ( 𝐵 gcd 𝐶 ) ) = ( 𝐴 gcd ( 𝐶 gcd 𝐵 ) ) )
4 gcdass ( ( 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ ∧ 𝐶 ∈ ℤ ) → ( ( 𝐴 gcd 𝐵 ) gcd 𝐶 ) = ( 𝐴 gcd ( 𝐵 gcd 𝐶 ) ) )
5 gcdass ( ( 𝐴 ∈ ℤ ∧ 𝐶 ∈ ℤ ∧ 𝐵 ∈ ℤ ) → ( ( 𝐴 gcd 𝐶 ) gcd 𝐵 ) = ( 𝐴 gcd ( 𝐶 gcd 𝐵 ) ) )
6 5 3com23 ( ( 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ ∧ 𝐶 ∈ ℤ ) → ( ( 𝐴 gcd 𝐶 ) gcd 𝐵 ) = ( 𝐴 gcd ( 𝐶 gcd 𝐵 ) ) )
7 3 4 6 3eqtr4d ( ( 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ ∧ 𝐶 ∈ ℤ ) → ( ( 𝐴 gcd 𝐵 ) gcd 𝐶 ) = ( ( 𝐴 gcd 𝐶 ) gcd 𝐵 ) )