HomeHome Intuitionistic Logic Explorer
Theorem List (p. 28 of 102)
< Previous  Next >
Bad symbols? Try the
GIF version.

Mirrors  >  Metamath Home Page  >  ILE Home Page  >  Theorem List Contents  >  Recent Proofs       This page: Page List

Theorem List for Intuitionistic Logic Explorer - 2701-2800   *Has distinct variable group(s)
TypeLabelDescription
Statement
 
Theoremralab 2701* Universal quantification over a class abstraction. (Contributed by Jeff Madsen, 10-Jun-2010.)
(𝑦 = 𝑥 → (𝜑𝜓))       (∀𝑥 ∈ {𝑦𝜑}𝜒 ↔ ∀𝑥(𝜓𝜒))
 
Theoremralrab 2702* Universal quantification over a restricted class abstraction. (Contributed by Jeff Madsen, 10-Jun-2010.)
(𝑦 = 𝑥 → (𝜑𝜓))       (∀𝑥 ∈ {𝑦𝐴𝜑}𝜒 ↔ ∀𝑥𝐴 (𝜓𝜒))
 
Theoremrexab 2703* Existential quantification over a class abstraction. (Contributed by Mario Carneiro, 23-Jan-2014.) (Revised by Mario Carneiro, 3-Sep-2015.)
(𝑦 = 𝑥 → (𝜑𝜓))       (∃𝑥 ∈ {𝑦𝜑}𝜒 ↔ ∃𝑥(𝜓𝜒))
 
Theoremrexrab 2704* Existential quantification over a class abstraction. (Contributed by Jeff Madsen, 17-Jun-2011.) (Revised by Mario Carneiro, 3-Sep-2015.)
(𝑦 = 𝑥 → (𝜑𝜓))       (∃𝑥 ∈ {𝑦𝐴𝜑}𝜒 ↔ ∃𝑥𝐴 (𝜓𝜒))
 
Theoremralab2 2705* Universal quantification over a class abstraction. (Contributed by Mario Carneiro, 3-Sep-2015.)
(𝑥 = 𝑦 → (𝜓𝜒))       (∀𝑥 ∈ {𝑦𝜑}𝜓 ↔ ∀𝑦(𝜑𝜒))
 
Theoremralrab2 2706* Universal quantification over a restricted class abstraction. (Contributed by Mario Carneiro, 3-Sep-2015.)
(𝑥 = 𝑦 → (𝜓𝜒))       (∀𝑥 ∈ {𝑦𝐴𝜑}𝜓 ↔ ∀𝑦𝐴 (𝜑𝜒))
 
Theoremrexab2 2707* Existential quantification over a class abstraction. (Contributed by Mario Carneiro, 3-Sep-2015.)
(𝑥 = 𝑦 → (𝜓𝜒))       (∃𝑥 ∈ {𝑦𝜑}𝜓 ↔ ∃𝑦(𝜑𝜒))
 
Theoremrexrab2 2708* Existential quantification over a class abstraction. (Contributed by Mario Carneiro, 3-Sep-2015.)
(𝑥 = 𝑦 → (𝜓𝜒))       (∃𝑥 ∈ {𝑦𝐴𝜑}𝜓 ↔ ∃𝑦𝐴 (𝜑𝜒))
 
Theoremabidnf 2709* Identity used to create closed-form versions of bound-variable hypothesis builders for class expressions. (Contributed by NM, 10-Nov-2005.) (Proof shortened by Mario Carneiro, 12-Oct-2016.)
(𝑥𝐴 → {𝑧 ∣ ∀𝑥 𝑧𝐴} = 𝐴)
 
Theoremdedhb 2710* A deduction theorem for converting the inference 𝑥𝐴 => 𝜑 into a closed theorem. Use nfa1 1434 and nfab 2182 to eliminate the hypothesis of the substitution instance 𝜓 of the inference. For converting the inference form into a deduction form, abidnf 2709 is useful. (Contributed by NM, 8-Dec-2006.)
(𝐴 = {𝑧 ∣ ∀𝑥 𝑧𝐴} → (𝜑𝜓))    &   𝜓       (𝑥𝐴𝜑)
 
Theoremeqeu 2711* A condition which implies existential uniqueness. (Contributed by Jeff Hankins, 8-Sep-2009.)
(𝑥 = 𝐴 → (𝜑𝜓))       ((𝐴𝐵𝜓 ∧ ∀𝑥(𝜑𝑥 = 𝐴)) → ∃!𝑥𝜑)
 
Theoremeueq 2712* Equality has existential uniqueness. (Contributed by NM, 25-Nov-1994.)
(𝐴 ∈ V ↔ ∃!𝑥 𝑥 = 𝐴)
 
Theoremeueq1 2713* Equality has existential uniqueness. (Contributed by NM, 5-Apr-1995.)
𝐴 ∈ V       ∃!𝑥 𝑥 = 𝐴
 
Theoremeueq2dc 2714* Equality has existential uniqueness (split into 2 cases). (Contributed by NM, 5-Apr-1995.)
𝐴 ∈ V    &   𝐵 ∈ V       (DECID 𝜑 → ∃!𝑥((𝜑𝑥 = 𝐴) ∨ (¬ 𝜑𝑥 = 𝐵)))
 
Theoremeueq3dc 2715* Equality has existential uniqueness (split into 3 cases). (Contributed by NM, 5-Apr-1995.) (Proof shortened by Mario Carneiro, 28-Sep-2015.)
𝐴 ∈ V    &   𝐵 ∈ V    &   𝐶 ∈ V    &    ¬ (𝜑𝜓)       (DECID 𝜑 → (DECID 𝜓 → ∃!𝑥((𝜑𝑥 = 𝐴) ∨ (¬ (𝜑𝜓) ∧ 𝑥 = 𝐵) ∨ (𝜓𝑥 = 𝐶))))
 
Theoremmoeq 2716* There is at most one set equal to a class. (Contributed by NM, 8-Mar-1995.)
∃*𝑥 𝑥 = 𝐴
 
Theoremmoeq3dc 2717* "At most one" property of equality (split into 3 cases). (Contributed by Jim Kingdon, 7-Jul-2018.)
𝐴 ∈ V    &   𝐵 ∈ V    &   𝐶 ∈ V    &    ¬ (𝜑𝜓)       (DECID 𝜑 → (DECID 𝜓 → ∃*𝑥((𝜑𝑥 = 𝐴) ∨ (¬ (𝜑𝜓) ∧ 𝑥 = 𝐵) ∨ (𝜓𝑥 = 𝐶))))
 
Theoremmosubt 2718* "At most one" remains true after substitution. (Contributed by Jim Kingdon, 18-Jan-2019.)
(∀𝑦∃*𝑥𝜑 → ∃*𝑥𝑦(𝑦 = 𝐴𝜑))
 
Theoremmosub 2719* "At most one" remains true after substitution. (Contributed by NM, 9-Mar-1995.)
∃*𝑥𝜑       ∃*𝑥𝑦(𝑦 = 𝐴𝜑)
 
Theoremmo2icl 2720* Theorem for inferring "at most one." (Contributed by NM, 17-Oct-1996.)
(∀𝑥(𝜑𝑥 = 𝐴) → ∃*𝑥𝜑)
 
Theoremmob2 2721* Consequence of "at most one." (Contributed by NM, 2-Jan-2015.)
(𝑥 = 𝐴 → (𝜑𝜓))       ((𝐴𝐵 ∧ ∃*𝑥𝜑𝜑) → (𝑥 = 𝐴𝜓))
 
Theoremmoi2 2722* Consequence of "at most one." (Contributed by NM, 29-Jun-2008.)
(𝑥 = 𝐴 → (𝜑𝜓))       (((𝐴𝐵 ∧ ∃*𝑥𝜑) ∧ (𝜑𝜓)) → 𝑥 = 𝐴)
 
Theoremmob 2723* Equality implied by "at most one." (Contributed by NM, 18-Feb-2006.)
(𝑥 = 𝐴 → (𝜑𝜓))    &   (𝑥 = 𝐵 → (𝜑𝜒))       (((𝐴𝐶𝐵𝐷) ∧ ∃*𝑥𝜑𝜓) → (𝐴 = 𝐵𝜒))
 
Theoremmoi 2724* Equality implied by "at most one." (Contributed by NM, 18-Feb-2006.)
(𝑥 = 𝐴 → (𝜑𝜓))    &   (𝑥 = 𝐵 → (𝜑𝜒))       (((𝐴𝐶𝐵𝐷) ∧ ∃*𝑥𝜑 ∧ (𝜓𝜒)) → 𝐴 = 𝐵)
 
Theoremmorex 2725* Derive membership from uniqueness. (Contributed by Jeff Madsen, 2-Sep-2009.)
𝐵 ∈ V    &   (𝑥 = 𝐵 → (𝜑𝜓))       ((∃𝑥𝐴 𝜑 ∧ ∃*𝑥𝜑) → (𝜓𝐵𝐴))
 
Theoremeuxfr2dc 2726* Transfer existential uniqueness from a variable 𝑥 to another variable 𝑦 contained in expression 𝐴. (Contributed by NM, 14-Nov-2004.)
𝐴 ∈ V    &   ∃*𝑦 𝑥 = 𝐴       (DECID𝑦𝑥(𝑥 = 𝐴𝜑) → (∃!𝑥𝑦(𝑥 = 𝐴𝜑) ↔ ∃!𝑦𝜑))
 
Theoremeuxfrdc 2727* Transfer existential uniqueness from a variable 𝑥 to another variable 𝑦 contained in expression 𝐴. (Contributed by NM, 14-Nov-2004.)
𝐴 ∈ V    &   ∃!𝑦 𝑥 = 𝐴    &   (𝑥 = 𝐴 → (𝜑𝜓))       (DECID𝑦𝑥(𝑥 = 𝐴𝜓) → (∃!𝑥𝜑 ↔ ∃!𝑦𝜓))
 
Theoremeuind 2728* Existential uniqueness via an indirect equality. (Contributed by NM, 11-Oct-2010.)
𝐵 ∈ V    &   (𝑥 = 𝑦 → (𝜑𝜓))    &   (𝑥 = 𝑦𝐴 = 𝐵)       ((∀𝑥𝑦((𝜑𝜓) → 𝐴 = 𝐵) ∧ ∃𝑥𝜑) → ∃!𝑧𝑥(𝜑𝑧 = 𝐴))
 
Theoremreu2 2729* A way to express restricted uniqueness. (Contributed by NM, 22-Nov-1994.)
(∃!𝑥𝐴 𝜑 ↔ (∃𝑥𝐴 𝜑 ∧ ∀𝑥𝐴𝑦𝐴 ((𝜑 ∧ [𝑦 / 𝑥]𝜑) → 𝑥 = 𝑦)))
 
Theoremreu6 2730* A way to express restricted uniqueness. (Contributed by NM, 20-Oct-2006.)
(∃!𝑥𝐴 𝜑 ↔ ∃𝑦𝐴𝑥𝐴 (𝜑𝑥 = 𝑦))
 
Theoremreu3 2731* A way to express restricted uniqueness. (Contributed by NM, 24-Oct-2006.)
(∃!𝑥𝐴 𝜑 ↔ (∃𝑥𝐴 𝜑 ∧ ∃𝑦𝐴𝑥𝐴 (𝜑𝑥 = 𝑦)))
 
Theoremreu6i 2732* A condition which implies existential uniqueness. (Contributed by Mario Carneiro, 2-Oct-2015.)
((𝐵𝐴 ∧ ∀𝑥𝐴 (𝜑𝑥 = 𝐵)) → ∃!𝑥𝐴 𝜑)
 
Theoremeqreu 2733* A condition which implies existential uniqueness. (Contributed by Mario Carneiro, 2-Oct-2015.)
(𝑥 = 𝐵 → (𝜑𝜓))       ((𝐵𝐴𝜓 ∧ ∀𝑥𝐴 (𝜑𝑥 = 𝐵)) → ∃!𝑥𝐴 𝜑)
 
Theoremrmo4 2734* Restricted "at most one" using implicit substitution. (Contributed by NM, 24-Oct-2006.) (Revised by NM, 16-Jun-2017.)
(𝑥 = 𝑦 → (𝜑𝜓))       (∃*𝑥𝐴 𝜑 ↔ ∀𝑥𝐴𝑦𝐴 ((𝜑𝜓) → 𝑥 = 𝑦))
 
Theoremreu4 2735* Restricted uniqueness using implicit substitution. (Contributed by NM, 23-Nov-1994.)
(𝑥 = 𝑦 → (𝜑𝜓))       (∃!𝑥𝐴 𝜑 ↔ (∃𝑥𝐴 𝜑 ∧ ∀𝑥𝐴𝑦𝐴 ((𝜑𝜓) → 𝑥 = 𝑦)))
 
Theoremreu7 2736* Restricted uniqueness using implicit substitution. (Contributed by NM, 24-Oct-2006.)
(𝑥 = 𝑦 → (𝜑𝜓))       (∃!𝑥𝐴 𝜑 ↔ (∃𝑥𝐴 𝜑 ∧ ∃𝑥𝐴𝑦𝐴 (𝜓𝑥 = 𝑦)))
 
Theoremreu8 2737* Restricted uniqueness using implicit substitution. (Contributed by NM, 24-Oct-2006.)
(𝑥 = 𝑦 → (𝜑𝜓))       (∃!𝑥𝐴 𝜑 ↔ ∃𝑥𝐴 (𝜑 ∧ ∀𝑦𝐴 (𝜓𝑥 = 𝑦)))
 
Theoremreueq 2738* Equality has existential uniqueness. (Contributed by Mario Carneiro, 1-Sep-2015.)
(𝐵𝐴 ↔ ∃!𝑥𝐴 𝑥 = 𝐵)
 
Theoremrmoan 2739 Restricted "at most one" still holds when a conjunct is added. (Contributed by NM, 16-Jun-2017.)
(∃*𝑥𝐴 𝜑 → ∃*𝑥𝐴 (𝜓𝜑))
 
Theoremrmoim 2740 Restricted "at most one" is preserved through implication (note wff reversal). (Contributed by Alexander van der Vekens, 17-Jun-2017.)
(∀𝑥𝐴 (𝜑𝜓) → (∃*𝑥𝐴 𝜓 → ∃*𝑥𝐴 𝜑))
 
Theoremrmoimia 2741 Restricted "at most one" is preserved through implication (note wff reversal). (Contributed by Alexander van der Vekens, 17-Jun-2017.)
(𝑥𝐴 → (𝜑𝜓))       (∃*𝑥𝐴 𝜓 → ∃*𝑥𝐴 𝜑)
 
Theoremrmoimi2 2742 Restricted "at most one" is preserved through implication (note wff reversal). (Contributed by Alexander van der Vekens, 17-Jun-2017.)
𝑥((𝑥𝐴𝜑) → (𝑥𝐵𝜓))       (∃*𝑥𝐵 𝜓 → ∃*𝑥𝐴 𝜑)
 
Theorem2reuswapdc 2743* A condition allowing swap of uniqueness and existential quantifiers. (Contributed by Thierry Arnoux, 7-Apr-2017.) (Revised by NM, 16-Jun-2017.)
(DECID𝑥𝑦(𝑥𝐴 ∧ (𝑦𝐵𝜑)) → (∀𝑥𝐴 ∃*𝑦𝐵 𝜑 → (∃!𝑥𝐴𝑦𝐵 𝜑 → ∃!𝑦𝐵𝑥𝐴 𝜑)))
 
Theoremreuind 2744* Existential uniqueness via an indirect equality. (Contributed by NM, 16-Oct-2010.)
(𝑥 = 𝑦 → (𝜑𝜓))    &   (𝑥 = 𝑦𝐴 = 𝐵)       ((∀𝑥𝑦(((𝐴𝐶𝜑) ∧ (𝐵𝐶𝜓)) → 𝐴 = 𝐵) ∧ ∃𝑥(𝐴𝐶𝜑)) → ∃!𝑧𝐶𝑥((𝐴𝐶𝜑) → 𝑧 = 𝐴))
 
Theorem2rmorex 2745* Double restricted quantification with "at most one," analogous to 2moex 1986. (Contributed by Alexander van der Vekens, 17-Jun-2017.)
(∃*𝑥𝐴𝑦𝐵 𝜑 → ∀𝑦𝐵 ∃*𝑥𝐴 𝜑)
 
Theoremnelrdva 2746* Deduce negative membership from an implication. (Contributed by Thierry Arnoux, 27-Nov-2017.)
((𝜑𝑥𝐴) → 𝑥𝐵)       (𝜑 → ¬ 𝐵𝐴)
 
2.1.7  Conditional equality (experimental)

This is a very useless definition, which "abbreviates" (𝑥 = 𝑦𝜑) as CondEq(𝑥 = 𝑦𝜑). What this display hides, though, is that the first expression, even though it has a shorter constant string, is actually much more complicated in its parse tree: it is parsed as (wi (wceq (cv vx) (cv vy)) wph), while the CondEq version is parsed as (wcdeq vx vy wph). It also allows us to give a name to the specific 3-ary operation (𝑥 = 𝑦𝜑).

This is all used as part of a metatheorem: we want to say that (𝑥 = 𝑦 → (𝜑(𝑥) ↔ 𝜑(𝑦))) and (𝑥 = 𝑦𝐴(𝑥) = 𝐴(𝑦)) are provable, for any expressions 𝜑(𝑥) or 𝐴(𝑥) in the language. The proof is by induction, so the base case is each of the primitives, which is why you will see a theorem for each of the set.mm primitive operations.

The metatheorem comes with a disjoint variables assumption: every variable in 𝜑(𝑥) is assumed disjoint from 𝑥 except 𝑥 itself. For such a proof by induction, we must consider each of the possible forms of 𝜑(𝑥). If it is a variable other than 𝑥, then we have CondEq(𝑥 = 𝑦𝐴 = 𝐴) or CondEq(𝑥 = 𝑦 → (𝜑𝜑)), which is provable by cdeqth 2751 and reflexivity. Since we are only working with class and wff expressions, it can't be 𝑥 itself in set.mm, but if it was we'd have to also prove CondEq(𝑥 = 𝑦𝑥 = 𝑦) (where set equality is being used on the right).

Otherwise, it is a primitive operation applied to smaller expressions. In these cases, for each setvar variable parameter to the operation, we must consider if it is equal to 𝑥 or not, which yields 2^n proof obligations. Luckily, all primitive operations in set.mm have either zero or one set variable, so we only need to prove one statement for the non-set constructors (like implication) and two for the constructors taking a set (the forall and the class builder).

In each of the primitive proofs, we are allowed to assume that 𝑦 is disjoint from 𝜑(𝑥) and vice versa, because this is maintained through the induction. This is how we satisfy the DV assumptions of cdeqab1 2756 and cdeqab 2754.

 
Syntaxwcdeq 2747 Extend wff notation to include conditional equality. This is a technical device used in the proof that is the not-free predicate, and that definitions are conservative as a result.
wff CondEq(𝑥 = 𝑦𝜑)
 
Definitiondf-cdeq 2748 Define conditional equality. All the notation to the left of the is fake; the parentheses and arrows are all part of the notation, which could equally well be written CondEq𝑥𝑦𝜑. On the right side is the actual implication arrow. The reason for this definition is to "flatten" the structure on the right side (whose tree structure is something like (wi (wceq (cv vx) (cv vy)) wph) ) into just (wcdeq vx vy wph). (Contributed by Mario Carneiro, 11-Aug-2016.)
(CondEq(𝑥 = 𝑦𝜑) ↔ (𝑥 = 𝑦𝜑))
 
Theoremcdeqi 2749 Deduce conditional equality. (Contributed by Mario Carneiro, 11-Aug-2016.)
(𝑥 = 𝑦𝜑)       CondEq(𝑥 = 𝑦𝜑)
 
Theoremcdeqri 2750 Property of conditional equality. (Contributed by Mario Carneiro, 11-Aug-2016.)
CondEq(𝑥 = 𝑦𝜑)       (𝑥 = 𝑦𝜑)
 
Theoremcdeqth 2751 Deduce conditional equality from a theorem. (Contributed by Mario Carneiro, 11-Aug-2016.)
𝜑       CondEq(𝑥 = 𝑦𝜑)
 
Theoremcdeqnot 2752 Distribute conditional equality over negation. (Contributed by Mario Carneiro, 11-Aug-2016.)
CondEq(𝑥 = 𝑦 → (𝜑𝜓))       CondEq(𝑥 = 𝑦 → (¬ 𝜑 ↔ ¬ 𝜓))
 
Theoremcdeqal 2753* Distribute conditional equality over quantification. (Contributed by Mario Carneiro, 11-Aug-2016.)
CondEq(𝑥 = 𝑦 → (𝜑𝜓))       CondEq(𝑥 = 𝑦 → (∀𝑧𝜑 ↔ ∀𝑧𝜓))
 
Theoremcdeqab 2754* Distribute conditional equality over abstraction. (Contributed by Mario Carneiro, 11-Aug-2016.)
CondEq(𝑥 = 𝑦 → (𝜑𝜓))       CondEq(𝑥 = 𝑦 → {𝑧𝜑} = {𝑧𝜓})
 
Theoremcdeqal1 2755* Distribute conditional equality over quantification. (Contributed by Mario Carneiro, 11-Aug-2016.)
CondEq(𝑥 = 𝑦 → (𝜑𝜓))       CondEq(𝑥 = 𝑦 → (∀𝑥𝜑 ↔ ∀𝑦𝜓))
 
Theoremcdeqab1 2756* Distribute conditional equality over abstraction. (Contributed by Mario Carneiro, 11-Aug-2016.)
CondEq(𝑥 = 𝑦 → (𝜑𝜓))       CondEq(𝑥 = 𝑦 → {𝑥𝜑} = {𝑦𝜓})
 
Theoremcdeqim 2757 Distribute conditional equality over implication. (Contributed by Mario Carneiro, 11-Aug-2016.)
CondEq(𝑥 = 𝑦 → (𝜑𝜓))    &   CondEq(𝑥 = 𝑦 → (𝜒𝜃))       CondEq(𝑥 = 𝑦 → ((𝜑𝜒) ↔ (𝜓𝜃)))
 
Theoremcdeqcv 2758 Conditional equality for set-to-class promotion. (Contributed by Mario Carneiro, 11-Aug-2016.)
CondEq(𝑥 = 𝑦𝑥 = 𝑦)
 
Theoremcdeqeq 2759 Distribute conditional equality over equality. (Contributed by Mario Carneiro, 11-Aug-2016.)
CondEq(𝑥 = 𝑦𝐴 = 𝐵)    &   CondEq(𝑥 = 𝑦𝐶 = 𝐷)       CondEq(𝑥 = 𝑦 → (𝐴 = 𝐶𝐵 = 𝐷))
 
Theoremcdeqel 2760 Distribute conditional equality over elementhood. (Contributed by Mario Carneiro, 11-Aug-2016.)
CondEq(𝑥 = 𝑦𝐴 = 𝐵)    &   CondEq(𝑥 = 𝑦𝐶 = 𝐷)       CondEq(𝑥 = 𝑦 → (𝐴𝐶𝐵𝐷))
 
Theoremnfcdeq 2761* If we have a conditional equality proof, where 𝜑 is 𝜑(𝑥) and 𝜓 is 𝜑(𝑦), and 𝜑(𝑥) in fact does not have 𝑥 free in it according to , then 𝜑(𝑥) ↔ 𝜑(𝑦) unconditionally. This proves that 𝑥𝜑 is actually a not-free predicate. (Contributed by Mario Carneiro, 11-Aug-2016.)
𝑥𝜑    &   CondEq(𝑥 = 𝑦 → (𝜑𝜓))       (𝜑𝜓)
 
Theoremnfccdeq 2762* Variation of nfcdeq 2761 for classes. (Contributed by Mario Carneiro, 11-Aug-2016.)
𝑥𝐴    &   CondEq(𝑥 = 𝑦𝐴 = 𝐵)       𝐴 = 𝐵
 
2.1.8  Russell's Paradox
 
Theoremru 2763 Russell's Paradox. Proposition 4.14 of [TakeutiZaring] p. 14.

In the late 1800s, Frege's Axiom of (unrestricted) Comprehension, expressed in our notation as 𝐴 ∈ V, asserted that any collection of sets 𝐴 is a set i.e. belongs to the universe V of all sets. In particular, by substituting {𝑥𝑥𝑥} (the "Russell class") for 𝐴, it asserted {𝑥𝑥𝑥} ∈ V, meaning that the "collection of all sets which are not members of themselves" is a set. However, here we prove {𝑥𝑥𝑥} ∉ V. This contradiction was discovered by Russell in 1901 (published in 1903), invalidating the Comprehension Axiom and leading to the collapse of Frege's system.

In 1908, Zermelo rectified this fatal flaw by replacing Comprehension with a weaker Subset (or Separation) Axiom asserting that 𝐴 is a set only when it is smaller than some other set 𝐵. The intuitionistic set theory IZF includes such a separation axiom, Axiom 6 of [Crosilla] p. "Axioms of CZF and IZF", which we include as ax-sep 3875. (Contributed by NM, 7-Aug-1994.)

{𝑥𝑥𝑥} ∉ V
 
2.1.9  Proper substitution of classes for sets
 
Syntaxwsbc 2764 Extend wff notation to include the proper substitution of a class for a set. Read this notation as "the proper substitution of class 𝐴 for setvar variable 𝑥 in wff 𝜑."
wff [𝐴 / 𝑥]𝜑
 
Definitiondf-sbc 2765 Define the proper substitution of a class for a set.

When 𝐴 is a proper class, our definition evaluates to false. This is somewhat arbitrary: we could have, instead, chosen the conclusion of sbc6 2789 for our definition, which always evaluates to true for proper classes.

Our definition also does not produce the same results as discussed in the proof of Theorem 6.6 of [Quine] p. 42 (although Theorem 6.6 itself does hold, as shown by dfsbcq 2766 below). Unfortunately, Quine's definition requires a recursive syntactical breakdown of 𝜑, and it does not seem possible to express it with a single closed formula.

If we did not want to commit to any specific proper class behavior, we could use this definition only to prove theorem dfsbcq 2766, which holds for both our definition and Quine's, and from which we can derive a weaker version of df-sbc 2765 in the form of sbc8g 2771. However, the behavior of Quine's definition at proper classes is similarly arbitrary, and for practical reasons (to avoid having to prove sethood of 𝐴 in every use of this definition) we allow direct reference to df-sbc 2765 and assert that [𝐴 / 𝑥]𝜑 is always false when 𝐴 is a proper class.

The related definition df-csb defines proper substitution into a class variable (as opposed to a wff variable). (Contributed by NM, 14-Apr-1995.) (Revised by NM, 25-Dec-2016.)

([𝐴 / 𝑥]𝜑𝐴 ∈ {𝑥𝜑})
 
Theoremdfsbcq 2766 This theorem, which is similar to Theorem 6.7 of [Quine] p. 42 and holds under both our definition and Quine's, provides us with a weak definition of the proper substitution of a class for a set. Since our df-sbc 2765 does not result in the same behavior as Quine's for proper classes, if we wished to avoid conflict with Quine's definition we could start with this theorem and dfsbcq2 2767 instead of df-sbc 2765. (dfsbcq2 2767 is needed because unlike Quine we do not overload the df-sb 1646 syntax.) As a consequence of these theorems, we can derive sbc8g 2771, which is a weaker version of df-sbc 2765 that leaves substitution undefined when 𝐴 is a proper class.

However, it is often a nuisance to have to prove the sethood hypothesis of sbc8g 2771, so we will allow direct use of df-sbc 2765. Proper substiution with a proper class is rarely needed, and when it is, we can simply use the expansion of Quine's definition. (Contributed by NM, 14-Apr-1995.)

(𝐴 = 𝐵 → ([𝐴 / 𝑥]𝜑[𝐵 / 𝑥]𝜑))
 
Theoremdfsbcq2 2767 This theorem, which is similar to Theorem 6.7 of [Quine] p. 42 and holds under both our definition and Quine's, relates logic substitution df-sb 1646 and substitution for class variables df-sbc 2765. Unlike Quine, we use a different syntax for each in order to avoid overloading it. See remarks in dfsbcq 2766. (Contributed by NM, 31-Dec-2016.)
(𝑦 = 𝐴 → ([𝑦 / 𝑥]𝜑[𝐴 / 𝑥]𝜑))
 
Theoremsbsbc 2768 Show that df-sb 1646 and df-sbc 2765 are equivalent when the class term 𝐴 in df-sbc 2765 is a setvar variable. This theorem lets us reuse theorems based on df-sb 1646 for proofs involving df-sbc 2765. (Contributed by NM, 31-Dec-2016.) (Proof modification is discouraged.)
([𝑦 / 𝑥]𝜑[𝑦 / 𝑥]𝜑)
 
Theoremsbceq1d 2769 Equality theorem for class substitution. (Contributed by Mario Carneiro, 9-Feb-2017.) (Revised by NM, 30-Jun-2018.)
(𝜑𝐴 = 𝐵)       (𝜑 → ([𝐴 / 𝑥]𝜓[𝐵 / 𝑥]𝜓))
 
Theoremsbceq1dd 2770 Equality theorem for class substitution. (Contributed by Mario Carneiro, 9-Feb-2017.) (Revised by NM, 30-Jun-2018.)
(𝜑𝐴 = 𝐵)    &   (𝜑[𝐴 / 𝑥]𝜓)       (𝜑[𝐵 / 𝑥]𝜓)
 
Theoremsbc8g 2771 This is the closest we can get to df-sbc 2765 if we start from dfsbcq 2766 (see its comments) and dfsbcq2 2767. (Contributed by NM, 18-Nov-2008.) (Proof shortened by Andrew Salmon, 29-Jun-2011.) (Proof modification is discouraged.)
(𝐴𝑉 → ([𝐴 / 𝑥]𝜑𝐴 ∈ {𝑥𝜑}))
 
Theoremsbcex 2772 By our definition of proper substitution, it can only be true if the substituted expression is a set. (Contributed by Mario Carneiro, 13-Oct-2016.)
([𝐴 / 𝑥]𝜑𝐴 ∈ V)
 
Theoremsbceq1a 2773 Equality theorem for class substitution. Class version of sbequ12 1654. (Contributed by NM, 26-Sep-2003.)
(𝑥 = 𝐴 → (𝜑[𝐴 / 𝑥]𝜑))
 
Theoremsbceq2a 2774 Equality theorem for class substitution. Class version of sbequ12r 1655. (Contributed by NM, 4-Jan-2017.)
(𝐴 = 𝑥 → ([𝐴 / 𝑥]𝜑𝜑))
 
Theoremspsbc 2775 Specialization: if a formula is true for all sets, it is true for any class which is a set. Similar to Theorem 6.11 of [Quine] p. 44. See also stdpc4 1658 and rspsbc 2840. (Contributed by NM, 16-Jan-2004.)
(𝐴𝑉 → (∀𝑥𝜑[𝐴 / 𝑥]𝜑))
 
Theoremspsbcd 2776 Specialization: if a formula is true for all sets, it is true for any class which is a set. Similar to Theorem 6.11 of [Quine] p. 44. See also stdpc4 1658 and rspsbc 2840. (Contributed by Mario Carneiro, 9-Feb-2017.)
(𝜑𝐴𝑉)    &   (𝜑 → ∀𝑥𝜓)       (𝜑[𝐴 / 𝑥]𝜓)
 
Theoremsbcth 2777 A substitution into a theorem remains true (when 𝐴 is a set). (Contributed by NM, 5-Nov-2005.)
𝜑       (𝐴𝑉[𝐴 / 𝑥]𝜑)
 
Theoremsbcthdv 2778* Deduction version of sbcth 2777. (Contributed by NM, 30-Nov-2005.) (Proof shortened by Andrew Salmon, 8-Jun-2011.)
(𝜑𝜓)       ((𝜑𝐴𝑉) → [𝐴 / 𝑥]𝜓)
 
Theoremsbcid 2779 An identity theorem for substitution. See sbid 1657. (Contributed by Mario Carneiro, 18-Feb-2017.)
([𝑥 / 𝑥]𝜑𝜑)
 
Theoremnfsbc1d 2780 Deduction version of nfsbc1 2781. (Contributed by NM, 23-May-2006.) (Revised by Mario Carneiro, 12-Oct-2016.)
(𝜑𝑥𝐴)       (𝜑 → Ⅎ𝑥[𝐴 / 𝑥]𝜓)
 
Theoremnfsbc1 2781 Bound-variable hypothesis builder for class substitution. (Contributed by Mario Carneiro, 12-Oct-2016.)
𝑥𝐴       𝑥[𝐴 / 𝑥]𝜑
 
Theoremnfsbc1v 2782* Bound-variable hypothesis builder for class substitution. (Contributed by Mario Carneiro, 12-Oct-2016.)
𝑥[𝐴 / 𝑥]𝜑
 
Theoremnfsbcd 2783 Deduction version of nfsbc 2784. (Contributed by NM, 23-Nov-2005.) (Revised by Mario Carneiro, 12-Oct-2016.)
𝑦𝜑    &   (𝜑𝑥𝐴)    &   (𝜑 → Ⅎ𝑥𝜓)       (𝜑 → Ⅎ𝑥[𝐴 / 𝑦]𝜓)
 
Theoremnfsbc 2784 Bound-variable hypothesis builder for class substitution. (Contributed by NM, 7-Sep-2014.) (Revised by Mario Carneiro, 12-Oct-2016.)
𝑥𝐴    &   𝑥𝜑       𝑥[𝐴 / 𝑦]𝜑
 
Theoremsbcco 2785* A composition law for class substitution. (Contributed by NM, 26-Sep-2003.) (Revised by Mario Carneiro, 13-Oct-2016.)
([𝐴 / 𝑦][𝑦 / 𝑥]𝜑[𝐴 / 𝑥]𝜑)
 
Theoremsbcco2 2786* A composition law for class substitution. Importantly, 𝑥 may occur free in the class expression substituted for 𝐴. (Contributed by NM, 5-Sep-2004.) (Proof shortened by Andrew Salmon, 8-Jun-2011.)
(𝑥 = 𝑦𝐴 = 𝐵)       ([𝑥 / 𝑦][𝐵 / 𝑥]𝜑[𝐴 / 𝑥]𝜑)
 
Theoremsbc5 2787* An equivalence for class substitution. (Contributed by NM, 23-Aug-1993.) (Revised by Mario Carneiro, 12-Oct-2016.)
([𝐴 / 𝑥]𝜑 ↔ ∃𝑥(𝑥 = 𝐴𝜑))
 
Theoremsbc6g 2788* An equivalence for class substitution. (Contributed by NM, 11-Oct-2004.) (Proof shortened by Andrew Salmon, 8-Jun-2011.)
(𝐴𝑉 → ([𝐴 / 𝑥]𝜑 ↔ ∀𝑥(𝑥 = 𝐴𝜑)))
 
Theoremsbc6 2789* An equivalence for class substitution. (Contributed by NM, 23-Aug-1993.) (Proof shortened by Eric Schmidt, 17-Jan-2007.)
𝐴 ∈ V       ([𝐴 / 𝑥]𝜑 ↔ ∀𝑥(𝑥 = 𝐴𝜑))
 
Theoremsbc7 2790* An equivalence for class substitution in the spirit of df-clab 2027. Note that 𝑥 and 𝐴 don't have to be distinct. (Contributed by NM, 18-Nov-2008.) (Revised by Mario Carneiro, 13-Oct-2016.)
([𝐴 / 𝑥]𝜑 ↔ ∃𝑦(𝑦 = 𝐴[𝑦 / 𝑥]𝜑))
 
Theoremcbvsbc 2791 Change bound variables in a wff substitution. (Contributed by Jeff Hankins, 19-Sep-2009.) (Proof shortened by Andrew Salmon, 8-Jun-2011.)
𝑦𝜑    &   𝑥𝜓    &   (𝑥 = 𝑦 → (𝜑𝜓))       ([𝐴 / 𝑥]𝜑[𝐴 / 𝑦]𝜓)
 
Theoremcbvsbcv 2792* Change the bound variable of a class substitution using implicit substitution. (Contributed by NM, 30-Sep-2008.) (Revised by Mario Carneiro, 13-Oct-2016.)
(𝑥 = 𝑦 → (𝜑𝜓))       ([𝐴 / 𝑥]𝜑[𝐴 / 𝑦]𝜓)
 
Theoremsbciegft 2793* Conversion of implicit substitution to explicit class substitution, using a bound-variable hypothesis instead of distinct variables. (Closed theorem version of sbciegf 2794.) (Contributed by NM, 10-Nov-2005.) (Revised by Mario Carneiro, 13-Oct-2016.)
((𝐴𝑉 ∧ Ⅎ𝑥𝜓 ∧ ∀𝑥(𝑥 = 𝐴 → (𝜑𝜓))) → ([𝐴 / 𝑥]𝜑𝜓))
 
Theoremsbciegf 2794* Conversion of implicit substitution to explicit class substitution. (Contributed by NM, 14-Dec-2005.) (Revised by Mario Carneiro, 13-Oct-2016.)
𝑥𝜓    &   (𝑥 = 𝐴 → (𝜑𝜓))       (𝐴𝑉 → ([𝐴 / 𝑥]𝜑𝜓))
 
Theoremsbcieg 2795* Conversion of implicit substitution to explicit class substitution. (Contributed by NM, 10-Nov-2005.)
(𝑥 = 𝐴 → (𝜑𝜓))       (𝐴𝑉 → ([𝐴 / 𝑥]𝜑𝜓))
 
Theoremsbcie2g 2796* Conversion of implicit substitution to explicit class substitution. This version of sbcie 2797 avoids a disjointness condition on 𝑥 and 𝐴 by substituting twice. (Contributed by Mario Carneiro, 15-Oct-2016.)
(𝑥 = 𝑦 → (𝜑𝜓))    &   (𝑦 = 𝐴 → (𝜓𝜒))       (𝐴𝑉 → ([𝐴 / 𝑥]𝜑𝜒))
 
Theoremsbcie 2797* Conversion of implicit substitution to explicit class substitution. (Contributed by NM, 4-Sep-2004.)
𝐴 ∈ V    &   (𝑥 = 𝐴 → (𝜑𝜓))       ([𝐴 / 𝑥]𝜑𝜓)
 
Theoremsbciedf 2798* Conversion of implicit substitution to explicit class substitution, deduction form. (Contributed by NM, 29-Dec-2014.)
(𝜑𝐴𝑉)    &   ((𝜑𝑥 = 𝐴) → (𝜓𝜒))    &   𝑥𝜑    &   (𝜑 → Ⅎ𝑥𝜒)       (𝜑 → ([𝐴 / 𝑥]𝜓𝜒))
 
Theoremsbcied 2799* Conversion of implicit substitution to explicit class substitution, deduction form. (Contributed by NM, 13-Dec-2014.)
(𝜑𝐴𝑉)    &   ((𝜑𝑥 = 𝐴) → (𝜓𝜒))       (𝜑 → ([𝐴 / 𝑥]𝜓𝜒))
 
Theoremsbcied2 2800* Conversion of implicit substitution to explicit class substitution, deduction form. (Contributed by NM, 13-Dec-2014.)
(𝜑𝐴𝑉)    &   (𝜑𝐴 = 𝐵)    &   ((𝜑𝑥 = 𝐵) → (𝜓𝜒))       (𝜑 → ([𝐴 / 𝑥]𝜓𝜒))
    < Previous  Next >

Page List
Jump to page: Contents  1 1-100 2 101-200 3 201-300 4 301-400 5 401-500 6 501-600 7 601-700 8 701-800 9 801-900 10 901-1000 11 1001-1100 12 1101-1200 13 1201-1300 14 1301-1400 15 1401-1500 16 1501-1600 17 1601-1700 18 1701-1800 19 1801-1900 20 1901-2000 21 2001-2100 22 2101-2200 23 2201-2300 24 2301-2400 25 2401-2500 26 2501-2600 27 2601-2700 28 2701-2800 29 2801-2900 30 2901-3000 31 3001-3100 32 3101-3200 33 3201-3300 34 3301-3400 35 3401-3500 36 3501-3600 37 3601-3700 38 3701-3800 39 3801-3900 40 3901-4000 41 4001-4100 42 4101-4200 43 4201-4300 44 4301-4400 45 4401-4500 46 4501-4600 47 4601-4700 48 4701-4800 49 4801-4900 50 4901-5000 51 5001-5100 52 5101-5200 53 5201-5300 54 5301-5400 55 5401-5500 56 5501-5600 57 5601-5700 58 5701-5800 59 5801-5900 60 5901-6000 61 6001-6100 62 6101-6200 63 6201-6300 64 6301-6400 65 6401-6500 66 6501-6600 67 6601-6700 68 6701-6800 69 6801-6900 70 6901-7000 71 7001-7100 72 7101-7200 73 7201-7300 74 7301-7400 75 7401-7500 76 7501-7600 77 7601-7700 78 7701-7800 79 7801-7900 80 7901-8000 81 8001-8100 82 8101-8200 83 8201-8300 84 8301-8400 85 8401-8500 86 8501-8600 87 8601-8700 88 8701-8800 89 8801-8900 90 8901-9000 91 9001-9100 92 9101-9200 93 9201-9300 94 9301-9400 95 9401-9500 96 9501-9600 97 9601-9700 98 9701-9800 99 9801-9900 100 9901-10000 101 10001-10100 102 10101-10124
  Copyright terms: Public domain < Previous  Next >