Home | New
Foundations Explorer Theorem List (p. 31 of 64) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > NFE Home Page > Theorem List Contents This page: Page List |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | rexrab 3001* | Existential quantification over a class abstraction. (Contributed by Jeff Madsen, 17-Jun-2011.) (Revised by Mario Carneiro, 3-Sep-2015.) |
⊢ (y = x → (φ ↔ ψ)) ⇒ ⊢ (∃x ∈ {y ∈ A ∣ φ}χ ↔ ∃x ∈ A (ψ ∧ χ)) | ||
Theorem | ralab2 3002* | Universal quantification over a class abstraction. (Contributed by Mario Carneiro, 3-Sep-2015.) |
⊢ (x = y → (ψ ↔ χ)) ⇒ ⊢ (∀x ∈ {y ∣ φ}ψ ↔ ∀y(φ → χ)) | ||
Theorem | ralrab2 3003* | Universal quantification over a restricted class abstraction. (Contributed by Mario Carneiro, 3-Sep-2015.) |
⊢ (x = y → (ψ ↔ χ)) ⇒ ⊢ (∀x ∈ {y ∈ A ∣ φ}ψ ↔ ∀y ∈ A (φ → χ)) | ||
Theorem | rexab2 3004* | Existential quantification over a class abstraction. (Contributed by Mario Carneiro, 3-Sep-2015.) |
⊢ (x = y → (ψ ↔ χ)) ⇒ ⊢ (∃x ∈ {y ∣ φ}ψ ↔ ∃y(φ ∧ χ)) | ||
Theorem | rexrab2 3005* | Existential quantification over a class abstraction. (Contributed by Mario Carneiro, 3-Sep-2015.) |
⊢ (x = y → (ψ ↔ χ)) ⇒ ⊢ (∃x ∈ {y ∈ A ∣ φ}ψ ↔ ∃y ∈ A (φ ∧ χ)) | ||
Theorem | abidnf 3006* | 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.) |
⊢ (ℲxA → {z ∣ ∀x z ∈ A} = A) | ||
Theorem | dedhb 3007* | A deduction theorem for converting the inference ⊢ ℲxA => ⊢ φ into a closed theorem. Use nfa1 1788 and nfab 2494 to eliminate the hypothesis of the substitution instance ψ of the inference. For converting the inference form into a deduction form, abidnf 3006 is useful. (Contributed by NM, 8-Dec-2006.) |
⊢ (A = {z ∣ ∀x z ∈ A} → (φ ↔ ψ)) & ⊢ ψ ⇒ ⊢ (ℲxA → φ) | ||
Theorem | eqeu 3008* | A condition which implies existential uniqueness. (Contributed by Jeff Hankins, 8-Sep-2009.) |
⊢ (x = A → (φ ↔ ψ)) ⇒ ⊢ ((A ∈ B ∧ ψ ∧ ∀x(φ → x = A)) → ∃!xφ) | ||
Theorem | eueq 3009* | Equality has existential uniqueness. (Contributed by NM, 25-Nov-1994.) |
⊢ (A ∈ V ↔ ∃!x x = A) | ||
Theorem | eueq1 3010* | Equality has existential uniqueness. (Contributed by NM, 5-Apr-1995.) |
⊢ A ∈ V ⇒ ⊢ ∃!x x = A | ||
Theorem | eueq2 3011* | Equality has existential uniqueness (split into 2 cases). (Contributed by NM, 5-Apr-1995.) |
⊢ A ∈ V & ⊢ B ∈ V ⇒ ⊢ ∃!x((φ ∧ x = A) ∨ (¬ φ ∧ x = B)) | ||
Theorem | eueq3 3012* | Equality has existential uniqueness (split into 3 cases). (Contributed by NM, 5-Apr-1995.) (Proof shortened by Mario Carneiro, 28-Sep-2015.) |
⊢ A ∈ V & ⊢ B ∈ V & ⊢ C ∈ V & ⊢ ¬ (φ ∧ ψ) ⇒ ⊢ ∃!x((φ ∧ x = A) ∨ (¬ (φ ∨ ψ) ∧ x = B) ∨ (ψ ∧ x = C)) | ||
Theorem | moeq 3013* | There is at most one set equal to a class. (Contributed by NM, 8-Mar-1995.) |
⊢ ∃*x x = A | ||
Theorem | moeq3 3014* | "At most one" property of equality (split into 3 cases). (The first 2 hypotheses could be eliminated with longer proof.) (Contributed by NM, 23-Apr-1995.) |
⊢ B ∈ V & ⊢ C ∈ V & ⊢ ¬ (φ ∧ ψ) ⇒ ⊢ ∃*x((φ ∧ x = A) ∨ (¬ (φ ∨ ψ) ∧ x = B) ∨ (ψ ∧ x = C)) | ||
Theorem | mosub 3015* | "At most one" remains true after substitution. (Contributed by NM, 9-Mar-1995.) |
⊢ ∃*xφ ⇒ ⊢ ∃*x∃y(y = A ∧ φ) | ||
Theorem | mo2icl 3016* | Theorem for inferring "at most one." (Contributed by NM, 17-Oct-1996.) |
⊢ (∀x(φ → x = A) → ∃*xφ) | ||
Theorem | mob2 3017* | Consequence of "at most one." (Contributed by NM, 2-Jan-2015.) |
⊢ (x = A → (φ ↔ ψ)) ⇒ ⊢ ((A ∈ B ∧ ∃*xφ ∧ φ) → (x = A ↔ ψ)) | ||
Theorem | moi2 3018* | Consequence of "at most one." (Contributed by NM, 29-Jun-2008.) |
⊢ (x = A → (φ ↔ ψ)) ⇒ ⊢ (((A ∈ B ∧ ∃*xφ) ∧ (φ ∧ ψ)) → x = A) | ||
Theorem | mob 3019* | Equality implied by "at most one." (Contributed by NM, 18-Feb-2006.) |
⊢ (x = A → (φ ↔ ψ)) & ⊢ (x = B → (φ ↔ χ)) ⇒ ⊢ (((A ∈ C ∧ B ∈ D) ∧ ∃*xφ ∧ ψ) → (A = B ↔ χ)) | ||
Theorem | moi 3020* | Equality implied by "at most one." (Contributed by NM, 18-Feb-2006.) |
⊢ (x = A → (φ ↔ ψ)) & ⊢ (x = B → (φ ↔ χ)) ⇒ ⊢ (((A ∈ C ∧ B ∈ D) ∧ ∃*xφ ∧ (ψ ∧ χ)) → A = B) | ||
Theorem | morex 3021* | Derive membership from uniqueness. (Contributed by Jeff Madsen, 2-Sep-2009.) |
⊢ B ∈ V & ⊢ (x = B → (φ ↔ ψ)) ⇒ ⊢ ((∃x ∈ A φ ∧ ∃*xφ) → (ψ → B ∈ A)) | ||
Theorem | euxfr2 3022* | Transfer existential uniqueness from a variable x to another variable y contained in expression A. (Contributed by NM, 14-Nov-2004.) |
⊢ A ∈ V & ⊢ ∃*y x = A ⇒ ⊢ (∃!x∃y(x = A ∧ φ) ↔ ∃!yφ) | ||
Theorem | euxfr 3023* | Transfer existential uniqueness from a variable x to another variable y contained in expression A. (Contributed by NM, 14-Nov-2004.) |
⊢ A ∈ V & ⊢ ∃!y x = A & ⊢ (x = A → (φ ↔ ψ)) ⇒ ⊢ (∃!xφ ↔ ∃!yψ) | ||
Theorem | euind 3024* | Existential uniqueness via an indirect equality. (Contributed by NM, 11-Oct-2010.) |
⊢ B ∈ V & ⊢ (x = y → (φ ↔ ψ)) & ⊢ (x = y → A = B) ⇒ ⊢ ((∀x∀y((φ ∧ ψ) → A = B) ∧ ∃xφ) → ∃!z∀x(φ → z = A)) | ||
Theorem | reu2 3025* | A way to express restricted uniqueness. (Contributed by NM, 22-Nov-1994.) |
⊢ (∃!x ∈ A φ ↔ (∃x ∈ A φ ∧ ∀x ∈ A ∀y ∈ A ((φ ∧ [y / x]φ) → x = y))) | ||
Theorem | reu6 3026* | A way to express restricted uniqueness. (Contributed by NM, 20-Oct-2006.) |
⊢ (∃!x ∈ A φ ↔ ∃y ∈ A ∀x ∈ A (φ ↔ x = y)) | ||
Theorem | reu3 3027* | A way to express restricted uniqueness. (Contributed by NM, 24-Oct-2006.) |
⊢ (∃!x ∈ A φ ↔ (∃x ∈ A φ ∧ ∃y ∈ A ∀x ∈ A (φ → x = y))) | ||
Theorem | reu6i 3028* | A condition which implies existential uniqueness. (Contributed by Mario Carneiro, 2-Oct-2015.) |
⊢ ((B ∈ A ∧ ∀x ∈ A (φ ↔ x = B)) → ∃!x ∈ A φ) | ||
Theorem | eqreu 3029* | A condition which implies existential uniqueness. (Contributed by Mario Carneiro, 2-Oct-2015.) |
⊢ (x = B → (φ ↔ ψ)) ⇒ ⊢ ((B ∈ A ∧ ψ ∧ ∀x ∈ A (φ → x = B)) → ∃!x ∈ A φ) | ||
Theorem | rmo4 3030* | Restricted "at most one" using implicit substitution. (Contributed by NM, 24-Oct-2006.) (Revised by NM, 16-Jun-2017.) |
⊢ (x = y → (φ ↔ ψ)) ⇒ ⊢ (∃*x ∈ A φ ↔ ∀x ∈ A ∀y ∈ A ((φ ∧ ψ) → x = y)) | ||
Theorem | reu4 3031* | Restricted uniqueness using implicit substitution. (Contributed by NM, 23-Nov-1994.) |
⊢ (x = y → (φ ↔ ψ)) ⇒ ⊢ (∃!x ∈ A φ ↔ (∃x ∈ A φ ∧ ∀x ∈ A ∀y ∈ A ((φ ∧ ψ) → x = y))) | ||
Theorem | reu7 3032* | Restricted uniqueness using implicit substitution. (Contributed by NM, 24-Oct-2006.) |
⊢ (x = y → (φ ↔ ψ)) ⇒ ⊢ (∃!x ∈ A φ ↔ (∃x ∈ A φ ∧ ∃x ∈ A ∀y ∈ A (ψ → x = y))) | ||
Theorem | reu8 3033* | Restricted unique existence using implicit substitution. (Contributed by NM, 24-Oct-2006.) |
⊢ (x = y → (φ ↔ ψ)) ⇒ ⊢ (∃!x ∈ A φ ↔ ∃x ∈ A (φ ∧ ∀y ∈ A (ψ → x = y))) | ||
Theorem | reueq 3034* | Equality has existential uniqueness. (Contributed by Mario Carneiro, 1-Sep-2015.) |
⊢ (B ∈ A ↔ ∃!x ∈ A x = B) | ||
Theorem | rmoan 3035 | Restricted "at most one" still holds when a conjunct is added. (Contributed by NM, 16-Jun-2017.) |
⊢ (∃*x ∈ A φ → ∃*x ∈ A (ψ ∧ φ)) | ||
Theorem | rmoim 3036 | Restricted "at most one" is preserved through implication (note wff reversal). (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ (∀x ∈ A (φ → ψ) → (∃*x ∈ A ψ → ∃*x ∈ A φ)) | ||
Theorem | rmoimia 3037 | Restricted "at most one" is preserved through implication (note wff reversal). (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ (x ∈ A → (φ → ψ)) ⇒ ⊢ (∃*x ∈ A ψ → ∃*x ∈ A φ) | ||
Theorem | rmoimi2 3038 | Restricted "at most one" is preserved through implication (note wff reversal). (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ ∀x((x ∈ A ∧ φ) → (x ∈ B ∧ ψ)) ⇒ ⊢ (∃*x ∈ B ψ → ∃*x ∈ A φ) | ||
Theorem | 2reuswap 3039* | A condition allowing swap of uniqueness and existential quantifiers. (Contributed by Thierry Arnoux, 7-Apr-2017.) (Revised by NM, 16-Jun-2017.) |
⊢ (∀x ∈ A ∃*y ∈ B φ → (∃!x ∈ A ∃y ∈ B φ → ∃!y ∈ B ∃x ∈ A φ)) | ||
Theorem | reuind 3040* | Existential uniqueness via an indirect equality. (Contributed by NM, 16-Oct-2010.) |
⊢ (x = y → (φ ↔ ψ)) & ⊢ (x = y → A = B) ⇒ ⊢ ((∀x∀y(((A ∈ C ∧ φ) ∧ (B ∈ C ∧ ψ)) → A = B) ∧ ∃x(A ∈ C ∧ φ)) → ∃!z ∈ C ∀x((A ∈ C ∧ φ) → z = A)) | ||
Theorem | 2rmorex 3041* | Double restricted quantification with "at most one," analogous to 2moex 2275. (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ (∃*x ∈ A ∃y ∈ B φ → ∀y ∈ B ∃*x ∈ A φ) | ||
Theorem | 2reu5lem1 3042* | Lemma for 2reu5 3045. Note that ∃!x ∈ A∃!y ∈ Bφ does not mean "there is exactly one x in A and exactly one y in B such that φ holds;" see comment for 2eu5 2288. (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ (∃!x ∈ A ∃!y ∈ B φ ↔ ∃!x∃!y(x ∈ A ∧ y ∈ B ∧ φ)) | ||
Theorem | 2reu5lem2 3043* | Lemma for 2reu5 3045. (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ (∀x ∈ A ∃*y ∈ B φ ↔ ∀x∃*y(x ∈ A ∧ y ∈ B ∧ φ)) | ||
Theorem | 2reu5lem3 3044* | Lemma for 2reu5 3045. This lemma is interesting in its own right, showing that existential restriction in the last conjunct (the "at most one" part) is optional; compare rmo2 3132. (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ ((∃!x ∈ A ∃!y ∈ B φ ∧ ∀x ∈ A ∃*y ∈ B φ) ↔ (∃x ∈ A ∃y ∈ B φ ∧ ∃z∃w∀x ∈ A ∀y ∈ B (φ → (x = z ∧ y = w)))) | ||
Theorem | 2reu5 3045* | Double restricted existential uniqueness in terms of restricted existential quantification and restricted universal quantification, analogous to 2eu5 2288 and reu3 3027. (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ ((∃!x ∈ A ∃!y ∈ B φ ∧ ∀x ∈ A ∃*y ∈ B φ) ↔ (∃x ∈ A ∃y ∈ B φ ∧ ∃z ∈ A ∃w ∈ B ∀x ∈ A ∀y ∈ B (φ → (x = z ∧ y = w)))) | ||
Theorem | ru 3046 |
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 A ∈ V, asserted that any collection of sets A is a set i.e. belongs to the universe V of all sets. In particular, by substituting {x ∣ x ∉ x} (the "Russell class") for A, it asserted {x ∣ x ∉ x} ∈ V, meaning that the "collection of all sets which are not members of themselves" is a set. However, here we prove {x ∣ x ∉ x} ∉ 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 ssex in set.mm asserting that A is a set only when it is smaller than some other set B. However, Zermelo was then faced with a "chicken and egg" problem of how to show B is a set, leading him to introduce the set-building axioms of Null Set 0ex 4111, Pairing prex 4113, Union uniex 4318, Power Set pwex 4330, and Infinity omex in set.mm to give him some starting sets to work with (all of which, before Russell's Paradox, were immediate consequences of Frege's Comprehension). In 1922 Fraenkel strengthened the Subset Axiom with our present Replacement Axiom funimaex in set.mm (whose modern formalization is due to Skolem, also in 1922). Thus, in a very real sense Russell's Paradox spawned the invention of ZF set theory and completely revised the foundations of mathematics! Another mainstream formalization of set theory, devised by von Neumann, Bernays, and Goedel, uses class variables rather than setvar variables as its primitives. The axiom system NBG in [Mendelson] p. 225 is suitable for a Metamath encoding. NBG is a conservative extension of ZF in that it proves exactly the same theorems as ZF that are expressible in the language of ZF. An advantage of NBG is that it is finitely axiomatizable - the Axiom of Replacement can be broken down into a finite set of formulas that eliminate its wff metavariable. Finite axiomatizability is required by some proof languages (although not by Metamath). There is a stronger version of NBG called Morse-Kelley (axiom system MK in [Mendelson] p. 287). Russell himself continued in a different direction, avoiding the paradox with his "theory of types." Quine extended Russell's ideas to formulate his New Foundations set theory (axiom system NF of [Quine] p. 331). In NF, the collection of all sets is a set, contradicting ZF and NBG set theories, and it has other bizarre consequences: when sets become too huge (beyond the size of those used in standard mathematics), the Axiom of Choice ac4 in set.mm and Cantor's Theorem canth in set.mm are provably false! (See ncanth in set.mm for some intuition behind the latter.) Recent results (as of 2014) seem to show that NF is equiconsistent to Z (ZF in which ax-sep in set.mm replaces ax-rep in set.mm) with ax-sep restricted to only bounded quantifiers. NF is finitely axiomatizable and can be encoded in Metamath using the axioms from T. Hailperin, "A set of axioms for logic," J. Symb. Logic 9:1-19 (1944). Under ZF set theory, every set is a member of the Russell class by elirrv in set.mm (derived from the Axiom of Regularity), so for us the Russell class equals the universe V (theorem ruv in set.mm). See ruALT in set.mm for an alternate proof of ru 3046 derived from that fact. (Contributed by NM, 7-Aug-1994.) |
⊢ {x ∣ x ∉ x} ∉ V | ||
Syntax | wsbc 3047 | Extend wff notation to include the proper substitution of a class for a set. Read this notation as "the proper substitution of class A for setvar variable x in wff φ." |
wff [̣A / x]̣φ | ||
Definition | df-sbc 3048 |
Define the proper substitution of a class for a set.
When A is a proper class, our definition evaluates to false. This is somewhat arbitrary: we could have, instead, chosen the conclusion of sbc6 3073 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 3049 below). For example, if A is a proper class, Quine's substitution of A for y in 0 ∈ y evaluates to 0 ∈ A rather than our falsehood. (This can be seen by substituting A, y, and 0 for alpha, beta, and gamma in Subcase 1 of Quine's discussion on p. 42.) 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 3049, which holds for both our definition and Quine's, and from which we can derive a weaker version of df-sbc 3048 in the form of sbc8g 3054. However, the behavior of Quine's definition at proper classes is similarly arbitrary, and for practical reasons (to avoid having to prove sethood of A in every use of this definition) we allow direct reference to df-sbc 3048 and assert that [̣A / x]̣φ is always false when A is a proper class. Theorem sbc2or 3055 shows the apparently "strongest" statement we can make regarding behavior at proper classes if we start from dfsbcq 3049. The related definition df-csb 3138 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.) |
⊢ ([̣A / x]̣φ ↔ A ∈ {x ∣ φ}) | ||
Theorem | dfsbcq 3049 |
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 3048 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 3050 instead of df-sbc 3048. (dfsbcq2 3050 is needed because
unlike Quine we do not overload the df-sb 1649 syntax.) As a consequence of
these theorems, we can derive sbc8g 3054, which is a weaker version of
df-sbc 3048 that leaves substitution undefined when A is a proper class.
However, it is often a nuisance to have to prove the sethood hypothesis of sbc8g 3054, so we will allow direct use of df-sbc 3048 after Theorem sbc2or 3055 below. 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.) |
⊢ (A = B → ([̣A / x]̣φ ↔ [̣B / x]̣φ)) | ||
Theorem | dfsbcq2 3050 | 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 1649 and substitution for class variables df-sbc 3048. Unlike Quine, we use a different syntax for each in order to avoid overloading it. See remarks in dfsbcq 3049. (Contributed by NM, 31-Dec-2016.) |
⊢ (y = A → ([y / x]φ ↔ [̣A / x]̣φ)) | ||
Theorem | sbsbc 3051 | Show that df-sb 1649 and df-sbc 3048 are equivalent when the class term A in df-sbc 3048 is a setvar variable. This theorem lets us reuse theorems based on df-sb 1649 for proofs involving df-sbc 3048. (Contributed by NM, 31-Dec-2016.) (Proof modification is discouraged.) |
⊢ ([y / x]φ ↔ [̣y / x]̣φ) | ||
Theorem | sbceq1d 3052 | Equality theorem for class substitution. (Contributed by Mario Carneiro, 9-Feb-2017.) |
⊢ (φ → A = B) ⇒ ⊢ (φ → ([̣A / x]̣φ ↔ [̣B / x]̣φ)) | ||
Theorem | sbceq1dd 3053 | Equality theorem for class substitution. (Contributed by Mario Carneiro, 9-Feb-2017.) |
⊢ (φ → A = B) & ⊢ (φ → [̣A / x]̣φ) ⇒ ⊢ (φ → [̣B / x]̣φ) | ||
Theorem | sbc8g 3054 | This is the closest we can get to df-sbc 3048 if we start from dfsbcq 3049 (see its comments) and dfsbcq2 3050. (Contributed by NM, 18-Nov-2008.) (Proof shortened by Andrew Salmon, 29-Jun-2011.) (Proof modification is discouraged.) |
⊢ (A ∈ V → ([̣A / x]̣φ ↔ A ∈ {x ∣ φ})) | ||
Theorem | sbc2or 3055* | The disjunction of two equivalences for class substitution does not require a class existence hypothesis. This theorem tells us that there are only 2 possibilities for [A / x]φ behavior at proper classes, matching the sbc5 3071 (false) and sbc6 3073 (true) conclusions. This is interesting since dfsbcq 3049 and dfsbcq2 3050 (from which it is derived) do not appear to say anything obvious about proper class behavior. Note that this theorem doesn't tell us that it is always one or the other at proper classes; it could "flip" between false (the first disjunct) and true (the second disjunct) as a function of some other variable y that φ or A may contain. (Contributed by NM, 11-Oct-2004.) (Proof modification is discouraged.) |
⊢ (([̣A / x]̣φ ↔ ∃x(x = A ∧ φ)) ∨ ([̣A / x]̣φ ↔ ∀x(x = A → φ))) | ||
Theorem | sbcex 3056 | 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.) |
⊢ ([̣A / x]̣φ → A ∈ V) | ||
Theorem | sbceq1a 3057 | Equality theorem for class substitution. Class version of sbequ12 1919. (Contributed by NM, 26-Sep-2003.) |
⊢ (x = A → (φ ↔ [̣A / x]̣φ)) | ||
Theorem | sbceq2a 3058 | Equality theorem for class substitution. Class version of sbequ12r 1920. (Contributed by NM, 4-Jan-2017.) |
⊢ (A = x → ([̣A / x]̣φ ↔ φ)) | ||
Theorem | spsbc 3059 | 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 2024 and rspsbc 3125. (Contributed by NM, 16-Jan-2004.) |
⊢ (A ∈ V → (∀xφ → [̣A / x]̣φ)) | ||
Theorem | spsbcd 3060 | 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 2024 and rspsbc 3125. (Contributed by Mario Carneiro, 9-Feb-2017.) |
⊢ (φ → A ∈ V) & ⊢ (φ → ∀xψ) ⇒ ⊢ (φ → [̣A / x]̣ψ) | ||
Theorem | sbcth 3061 | A substitution into a theorem remains true (when A is a set). (Contributed by NM, 5-Nov-2005.) |
⊢ φ ⇒ ⊢ (A ∈ V → [̣A / x]̣φ) | ||
Theorem | sbcthdv 3062* | Deduction version of sbcth 3061. (Contributed by NM, 30-Nov-2005.) (Proof shortened by Andrew Salmon, 8-Jun-2011.) |
⊢ (φ → ψ) ⇒ ⊢ ((φ ∧ A ∈ V) → [̣A / x]̣ψ) | ||
Theorem | sbcid 3063 | An identity theorem for substitution. See sbid 1922. (Contributed by Mario Carneiro, 18-Feb-2017.) |
⊢ ([̣x / x]̣φ ↔ φ) | ||
Theorem | nfsbc1d 3064 | Deduction version of nfsbc1 3065. (Contributed by NM, 23-May-2006.) (Revised by Mario Carneiro, 12-Oct-2016.) |
⊢ (φ → ℲxA) ⇒ ⊢ (φ → Ⅎx[̣A / x]̣ψ) | ||
Theorem | nfsbc1 3065 | Bound-variable hypothesis builder for class substitution. (Contributed by Mario Carneiro, 12-Oct-2016.) |
⊢ ℲxA ⇒ ⊢ Ⅎx[̣A / x]̣φ | ||
Theorem | nfsbc1v 3066* | Bound-variable hypothesis builder for class substitution. (Contributed by Mario Carneiro, 12-Oct-2016.) |
⊢ Ⅎx[̣A / x]̣φ | ||
Theorem | nfsbcd 3067 | Deduction version of nfsbc 3068. (Contributed by NM, 23-Nov-2005.) (Revised by Mario Carneiro, 12-Oct-2016.) |
⊢ Ⅎyφ & ⊢ (φ → ℲxA) & ⊢ (φ → Ⅎxψ) ⇒ ⊢ (φ → Ⅎx[̣A / y]̣ψ) | ||
Theorem | nfsbc 3068 | Bound-variable hypothesis builder for class substitution. (Contributed by NM, 7-Sep-2014.) (Revised by Mario Carneiro, 12-Oct-2016.) |
⊢ ℲxA & ⊢ Ⅎxφ ⇒ ⊢ Ⅎx[̣A / y]̣φ | ||
Theorem | sbcco 3069* | A composition law for class substitution. (Contributed by NM, 26-Sep-2003.) (Revised by Mario Carneiro, 13-Oct-2016.) |
⊢ ([̣A / y]̣[̣y / x]̣φ ↔ [̣A / x]̣φ) | ||
Theorem | sbcco2 3070* | A composition law for class substitution. Importantly, x may occur free in the class expression substituted for A. (Contributed by NM, 5-Sep-2004.) (Proof shortened by Andrew Salmon, 8-Jun-2011.) |
⊢ (x = y → A = B) ⇒ ⊢ ([̣x / y]̣[̣B / x]̣φ ↔ [̣A / x]̣φ) | ||
Theorem | sbc5 3071* | An equivalence for class substitution. (Contributed by NM, 23-Aug-1993.) (Revised by Mario Carneiro, 12-Oct-2016.) |
⊢ ([̣A / x]̣φ ↔ ∃x(x = A ∧ φ)) | ||
Theorem | sbc6g 3072* | An equivalence for class substitution. (Contributed by NM, 11-Oct-2004.) (Proof shortened by Andrew Salmon, 8-Jun-2011.) |
⊢ (A ∈ V → ([̣A / x]̣φ ↔ ∀x(x = A → φ))) | ||
Theorem | sbc6 3073* | An equivalence for class substitution. (Contributed by NM, 23-Aug-1993.) (Proof shortened by Eric Schmidt, 17-Jan-2007.) |
⊢ A ∈ V ⇒ ⊢ ([̣A / x]̣φ ↔ ∀x(x = A → φ)) | ||
Theorem | sbc7 3074* | An equivalence for class substitution in the spirit of df-clab 2340. Note that x and A don't have to be distinct. (Contributed by NM, 18-Nov-2008.) (Revised by Mario Carneiro, 13-Oct-2016.) |
⊢ ([̣A / x]̣φ ↔ ∃y(y = A ∧ [̣y / x]̣φ)) | ||
Theorem | cbvsbc 3075 | Change bound variables in a wff substitution. (Contributed by Jeff Hankins, 19-Sep-2009.) (Proof shortened by Andrew Salmon, 8-Jun-2011.) |
⊢ Ⅎyφ & ⊢ Ⅎxψ & ⊢ (x = y → (φ ↔ ψ)) ⇒ ⊢ ([̣A / x]̣φ ↔ [̣A / y]̣ψ) | ||
Theorem | cbvsbcv 3076* | Change the bound variable of a class substitution using implicit substitution. (Contributed by NM, 30-Sep-2008.) (Revised by Mario Carneiro, 13-Oct-2016.) |
⊢ (x = y → (φ ↔ ψ)) ⇒ ⊢ ([̣A / x]̣φ ↔ [̣A / y]̣ψ) | ||
Theorem | sbciegft 3077* | Conversion of implicit substitution to explicit class substitution, using a bound-variable hypothesis instead of distinct variables. (Closed theorem version of sbciegf 3078.) (Contributed by NM, 10-Nov-2005.) (Revised by Mario Carneiro, 13-Oct-2016.) |
⊢ ((A ∈ V ∧ Ⅎxψ ∧ ∀x(x = A → (φ ↔ ψ))) → ([̣A / x]̣φ ↔ ψ)) | ||
Theorem | sbciegf 3078* | Conversion of implicit substitution to explicit class substitution. (Contributed by NM, 14-Dec-2005.) (Revised by Mario Carneiro, 13-Oct-2016.) |
⊢ Ⅎxψ & ⊢ (x = A → (φ ↔ ψ)) ⇒ ⊢ (A ∈ V → ([̣A / x]̣φ ↔ ψ)) | ||
Theorem | sbcieg 3079* | Conversion of implicit substitution to explicit class substitution. (Contributed by NM, 10-Nov-2005.) |
⊢ (x = A → (φ ↔ ψ)) ⇒ ⊢ (A ∈ V → ([̣A / x]̣φ ↔ ψ)) | ||
Theorem | sbcie2g 3080* | Conversion of implicit substitution to explicit class substitution. This version of sbcie 3081 avoids a disjointness condition on x, A by substituting twice. (Contributed by Mario Carneiro, 15-Oct-2016.) |
⊢ (x = y → (φ ↔ ψ)) & ⊢ (y = A → (ψ ↔ χ)) ⇒ ⊢ (A ∈ V → ([̣A / x]̣φ ↔ χ)) | ||
Theorem | sbcie 3081* | Conversion of implicit substitution to explicit class substitution. (Contributed by NM, 4-Sep-2004.) |
⊢ A ∈ V & ⊢ (x = A → (φ ↔ ψ)) ⇒ ⊢ ([̣A / x]̣φ ↔ ψ) | ||
Theorem | sbciedf 3082* | Conversion of implicit substitution to explicit class substitution, deduction form. (Contributed by NM, 29-Dec-2014.) |
⊢ (φ → A ∈ V) & ⊢ ((φ ∧ x = A) → (ψ ↔ χ)) & ⊢ Ⅎxφ & ⊢ (φ → Ⅎxχ) ⇒ ⊢ (φ → ([̣A / x]̣ψ ↔ χ)) | ||
Theorem | sbcied 3083* | Conversion of implicit substitution to explicit class substitution, deduction form. (Contributed by NM, 13-Dec-2014.) |
⊢ (φ → A ∈ V) & ⊢ ((φ ∧ x = A) → (ψ ↔ χ)) ⇒ ⊢ (φ → ([̣A / x]̣ψ ↔ χ)) | ||
Theorem | sbcied2 3084* | Conversion of implicit substitution to explicit class substitution, deduction form. (Contributed by NM, 13-Dec-2014.) |
⊢ (φ → A ∈ V) & ⊢ (φ → A = B) & ⊢ ((φ ∧ x = B) → (ψ ↔ χ)) ⇒ ⊢ (φ → ([̣A / x]̣ψ ↔ χ)) | ||
Theorem | elrabsf 3085 | Membership in a restricted class abstraction, expressed with explicit class substitution. (The variation elrabf 2994 has implicit substitution). The hypothesis specifies that x must not be a free variable in B. (Contributed by NM, 30-Sep-2003.) (Proof shortened by Mario Carneiro, 13-Oct-2016.) |
⊢ ℲxB ⇒ ⊢ (A ∈ {x ∈ B ∣ φ} ↔ (A ∈ B ∧ [̣A / x]̣φ)) | ||
Theorem | eqsbc1 3086* | Substitution for the left-hand side in an equality. Class version of eqsb1 2454. (Contributed by Andrew Salmon, 29-Jun-2011.) |
⊢ (A ∈ V → ([̣A / x]̣x = B ↔ A = B)) | ||
Theorem | sbcng 3087 | Move negation in and out of class substitution. (Contributed by NM, 16-Jan-2004.) |
⊢ (A ∈ V → ([̣A / x]̣ ¬ φ ↔ ¬ [̣A / x]̣φ)) | ||
Theorem | sbcimg 3088 | Distribution of class substitution over implication. (Contributed by NM, 16-Jan-2004.) |
⊢ (A ∈ V → ([̣A / x]̣(φ → ψ) ↔ ([̣A / x]̣φ → [̣A / x]̣ψ))) | ||
Theorem | sbcan 3089 | Distribution of class substitution over conjunction. (Contributed by NM, 31-Dec-2016.) |
⊢ ([̣A / x]̣(φ ∧ ψ) ↔ ([̣A / x]̣φ ∧ [̣A / x]̣ψ)) | ||
Theorem | sbcang 3090 | Distribution of class substitution over conjunction. (Contributed by NM, 21-May-2004.) |
⊢ (A ∈ V → ([̣A / x]̣(φ ∧ ψ) ↔ ([̣A / x]̣φ ∧ [̣A / x]̣ψ))) | ||
Theorem | sbcor 3091 | Distribution of class substitution over disjunction. (Contributed by NM, 31-Dec-2016.) |
⊢ ([̣A / x]̣(φ ∨ ψ) ↔ ([̣A / x]̣φ ∨ [̣A / x]̣ψ)) | ||
Theorem | sbcorg 3092 | Distribution of class substitution over disjunction. (Contributed by NM, 21-May-2004.) |
⊢ (A ∈ V → ([̣A / x]̣(φ ∨ ψ) ↔ ([̣A / x]̣φ ∨ [̣A / x]̣ψ))) | ||
Theorem | sbcbig 3093 | Distribution of class substitution over biconditional. (Contributed by Raph Levien, 10-Apr-2004.) |
⊢ (A ∈ V → ([̣A / x]̣(φ ↔ ψ) ↔ ([̣A / x]̣φ ↔ [̣A / x]̣ψ))) | ||
Theorem | sbcal 3094* | Move universal quantifier in and out of class substitution. (Contributed by NM, 31-Dec-2016.) |
⊢ ([̣A / y]̣∀xφ ↔ ∀x[̣A / y]̣φ) | ||
Theorem | sbcalg 3095* | Move universal quantifier in and out of class substitution. (Contributed by NM, 16-Jan-2004.) |
⊢ (A ∈ V → ([̣A / y]̣∀xφ ↔ ∀x[̣A / y]̣φ)) | ||
Theorem | sbcex2 3096* | Move existential quantifier in and out of class substitution. (Contributed by NM, 21-May-2004.) |
⊢ ([̣A / y]̣∃xφ ↔ ∃x[̣A / y]̣φ) | ||
Theorem | sbcexg 3097* | Move existential quantifier in and out of class substitution. (Contributed by NM, 21-May-2004.) |
⊢ (A ∈ V → ([̣A / y]̣∃xφ ↔ ∃x[̣A / y]̣φ)) | ||
Theorem | sbceqal 3098* | Set theory version of sbeqal1 in set.mm. (Contributed by Andrew Salmon, 28-Jun-2011.) |
⊢ (A ∈ V → (∀x(x = A → x = B) → A = B)) | ||
Theorem | sbeqalb 3099* | Theorem *14.121 in [WhiteheadRussell] p. 185. (Contributed by Andrew Salmon, 28-Jun-2011.) (Proof shortened by Wolf Lammen, 9-May-2013.) |
⊢ (A ∈ V → ((∀x(φ ↔ x = A) ∧ ∀x(φ ↔ x = B)) → A = B)) | ||
Theorem | sbcbid 3100 | Formula-building deduction rule for class substitution. (Contributed by NM, 29-Dec-2014.) |
⊢ Ⅎxφ & ⊢ (φ → (ψ ↔ χ)) ⇒ ⊢ (φ → ([̣A / x]̣ψ ↔ [̣A / x]̣χ)) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |