Metamath Proof Explorer


Theorem stdpc4

Description: The specialization axiom of standard predicate calculus. It states that if a statement ph holds for all x , then it also holds for the specific case of t (properly) substituted for x . Translated to traditional notation, it can be read: " A. x ph ( x ) -> ph ( t ) , provided that t is free for x in ph ( x ) ". Axiom 4 of Mendelson p. 69. See also spsbc and rspsbc . (Contributed by NM, 14-May-1993) Revise df-sb . (Revised by BJ, 22-Dec-2020)

Ref Expression
Assertion stdpc4 ( ∀ 𝑥 𝜑 → [ 𝑡 / 𝑥 ] 𝜑 )

Proof

Step Hyp Ref Expression
1 ala1 ( ∀ 𝑥 𝜑 → ∀ 𝑥 ( 𝑥 = 𝑦𝜑 ) )
2 1 a1d ( ∀ 𝑥 𝜑 → ( 𝑦 = 𝑡 → ∀ 𝑥 ( 𝑥 = 𝑦𝜑 ) ) )
3 2 alrimiv ( ∀ 𝑥 𝜑 → ∀ 𝑦 ( 𝑦 = 𝑡 → ∀ 𝑥 ( 𝑥 = 𝑦𝜑 ) ) )
4 df-sb ( [ 𝑡 / 𝑥 ] 𝜑 ↔ ∀ 𝑦 ( 𝑦 = 𝑡 → ∀ 𝑥 ( 𝑥 = 𝑦𝜑 ) ) )
5 3 4 sylibr ( ∀ 𝑥 𝜑 → [ 𝑡 / 𝑥 ] 𝜑 )