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

Axiom ax-13 2234
Description: Axiom of Quantified Equality. One of the equality and substitution axioms of predicate calculus with equality.

An equivalent way to express this axiom that may be easier to understand is 𝑥 = 𝑦 → (¬ 𝑥 = 𝑧 → (𝑦 = 𝑧 → ∀𝑥𝑦 = 𝑧))) (see ax13b 1951). Recall that in the intended interpretation, our variables are metavariables ranging over the variables of predicate calculus (the object language). In order for the first antecedent ¬ 𝑥 = 𝑦 to hold, 𝑥 and 𝑦 must have different values and thus cannot be the same object-language variable (so they are effectively "distinct variables" even though no $d is present). Similarly, 𝑥 and 𝑧 cannot be the same object-language variable. Therefore, 𝑥 will not occur in the wff 𝑦 = 𝑧 when the first two antecedents hold, so analogous to ax-5 1827, the conclusion (𝑦 = 𝑧 → ∀𝑥𝑦 = 𝑧) follows. Note that ax-5 1827 cannot prove this because its distinct variable ($d) requirement is not satisfied directly but only indirectly (outside of Metamath) by the argument above.

The original version of this axiom was ax-c9 33193 and was replaced with this shorter ax-13 2234 in December 2015. The old axiom is proved from this one as theorem axc9 2290.

The primary purpose of this axiom is to provide a way to introduce the quantifier 𝑥 on 𝑦 = 𝑧 even when 𝑥 and 𝑦 are substituted with the same variable. In this case, the first antecedent becomes ¬ 𝑥 = 𝑥 and the axiom still holds.

Although this version is shorter, the original version axc9 2290 may be more practical to work with because of the "distinctor" form of its antecedents. A typical application of axc9 2290 is in dvelimh 2324 which converts a distinct variable pair to the distinctor antecedent ¬ ∀𝑥𝑥 = 𝑦. In particular, it is conjectured that it is not possible to prove ax6 2239 from ax6v 1876 without this axiom.

This axiom can be weakened if desired by adding distinct variable restrictions on pairs 𝑥, 𝑧 and 𝑦, 𝑧. To show that, we add these restrictions to theorem ax13v 2235 and use only ax13v 2235 for further derivations. Thus, ax13v 2235 should be the only theorem referencing this axiom. Other theorems can reference either ax13v 2235 (preferred) or ax13 2237 (if the stronger form is needed).

This axiom scheme is logically redundant (see ax13w 2000) but is used as an auxiliary axiom scheme to achieve scheme completeness (i.e. so that all possible cases of bundling can be proved; see text linked at mmtheorems.html#ax6dgen). It is not known whether this axiom can be derived from the others. (Contributed by NM, 21-Dec-2015.) (New usage is discouraged.)

Assertion
Ref Expression
ax-13 𝑥 = 𝑦 → (𝑦 = 𝑧 → ∀𝑥 𝑦 = 𝑧))

Detailed syntax breakdown of Axiom ax-13
StepHypRef Expression
1 vx . . . 4 setvar 𝑥
2 vy . . . 4 setvar 𝑦
31, 2weq 1861 . . 3 wff 𝑥 = 𝑦
43wn 3 . 2 wff ¬ 𝑥 = 𝑦
5 vz . . . 4 setvar 𝑧
62, 5weq 1861 . . 3 wff 𝑦 = 𝑧
76, 1wal 1473 . . 3 wff 𝑥 𝑦 = 𝑧
86, 7wi 4 . 2 wff (𝑦 = 𝑧 → ∀𝑥 𝑦 = 𝑧)
94, 8wi 4 1 wff 𝑥 = 𝑦 → (𝑦 = 𝑧 → ∀𝑥 𝑦 = 𝑧))
Colors of variables: wff setvar class
This axiom is referenced by:  ax13v  2235
  Copyright terms: Public domain W3C validator