Metamath Proof Explorer


Theorem axc11n

Description: Derive set.mm's original ax-c11n from others. Commutation law for identical variable specifiers. The antecedent and consequent are true when x and y are substituted with the same variable. Lemma L12 in Megill p. 445 (p. 12 of the preprint). If a disjoint variable condition is added on x and y , then this becomes an instance of aevlem . Use aecom instead when this does not lengthen the proof. Usage of this theorem is discouraged because it depends on ax-13 . (Contributed by NM, 10-May-1993) (Revised by NM, 7-Nov-2015) (Proof shortened by Wolf Lammen, 6-Mar-2018) (Revised by Wolf Lammen, 30-Nov-2019) (Proof shortened by BJ, 29-Mar-2021) (Proof shortened by Wolf Lammen, 2-Jul-2021) (New usage is discouraged.)

Ref Expression
Assertion axc11n ( ∀ 𝑥 𝑥 = 𝑦 → ∀ 𝑦 𝑦 = 𝑥 )

Proof

Step Hyp Ref Expression
1 dveeq1 ( ¬ ∀ 𝑦 𝑦 = 𝑥 → ( 𝑥 = 𝑧 → ∀ 𝑦 𝑥 = 𝑧 ) )
2 1 com12 ( 𝑥 = 𝑧 → ( ¬ ∀ 𝑦 𝑦 = 𝑥 → ∀ 𝑦 𝑥 = 𝑧 ) )
3 axc11r ( ∀ 𝑥 𝑥 = 𝑦 → ( ∀ 𝑦 𝑥 = 𝑧 → ∀ 𝑥 𝑥 = 𝑧 ) )
4 aev ( ∀ 𝑥 𝑥 = 𝑧 → ∀ 𝑦 𝑦 = 𝑥 )
5 3 4 syl6 ( ∀ 𝑥 𝑥 = 𝑦 → ( ∀ 𝑦 𝑥 = 𝑧 → ∀ 𝑦 𝑦 = 𝑥 ) )
6 2 5 syl9 ( 𝑥 = 𝑧 → ( ∀ 𝑥 𝑥 = 𝑦 → ( ¬ ∀ 𝑦 𝑦 = 𝑥 → ∀ 𝑦 𝑦 = 𝑥 ) ) )
7 ax6evr 𝑧 𝑥 = 𝑧
8 6 7 exlimiiv ( ∀ 𝑥 𝑥 = 𝑦 → ( ¬ ∀ 𝑦 𝑦 = 𝑥 → ∀ 𝑦 𝑦 = 𝑥 ) )
9 8 pm2.18d ( ∀ 𝑥 𝑥 = 𝑦 → ∀ 𝑦 𝑦 = 𝑥 )