![]() |
Metamath
Proof Explorer Theorem List (p. 35 of 429) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > MPE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Color key: | ![]() (1-27903) |
![]() (27904-29428) |
![]() (29429-42879) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | ralrab 3401* | Universal quantification over a restricted class abstraction. (Contributed by Jeff Madsen, 10-Jun-2010.) |
⊢ (𝑦 = 𝑥 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 ∈ {𝑦 ∈ 𝐴 ∣ 𝜑}𝜒 ↔ ∀𝑥 ∈ 𝐴 (𝜓 → 𝜒)) | ||
Theorem | rexab 3402* | Existential quantification over a class abstraction. (Contributed by Mario Carneiro, 23-Jan-2014.) (Revised by Mario Carneiro, 3-Sep-2015.) |
⊢ (𝑦 = 𝑥 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃𝑥 ∈ {𝑦 ∣ 𝜑}𝜒 ↔ ∃𝑥(𝜓 ∧ 𝜒)) | ||
Theorem | rexrab 3403* | Existential quantification over a class abstraction. (Contributed by Jeff Madsen, 17-Jun-2011.) (Revised by Mario Carneiro, 3-Sep-2015.) |
⊢ (𝑦 = 𝑥 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃𝑥 ∈ {𝑦 ∈ 𝐴 ∣ 𝜑}𝜒 ↔ ∃𝑥 ∈ 𝐴 (𝜓 ∧ 𝜒)) | ||
Theorem | ralab2 3404* | Universal quantification over a class abstraction. (Contributed by Mario Carneiro, 3-Sep-2015.) |
⊢ (𝑥 = 𝑦 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (∀𝑥 ∈ {𝑦 ∣ 𝜑}𝜓 ↔ ∀𝑦(𝜑 → 𝜒)) | ||
Theorem | ralrab2 3405* | Universal quantification over a restricted class abstraction. (Contributed by Mario Carneiro, 3-Sep-2015.) |
⊢ (𝑥 = 𝑦 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (∀𝑥 ∈ {𝑦 ∈ 𝐴 ∣ 𝜑}𝜓 ↔ ∀𝑦 ∈ 𝐴 (𝜑 → 𝜒)) | ||
Theorem | rexab2 3406* | Existential quantification over a class abstraction. (Contributed by Mario Carneiro, 3-Sep-2015.) |
⊢ (𝑥 = 𝑦 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (∃𝑥 ∈ {𝑦 ∣ 𝜑}𝜓 ↔ ∃𝑦(𝜑 ∧ 𝜒)) | ||
Theorem | rexrab2 3407* | Existential quantification over a class abstraction. (Contributed by Mario Carneiro, 3-Sep-2015.) |
⊢ (𝑥 = 𝑦 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (∃𝑥 ∈ {𝑦 ∈ 𝐴 ∣ 𝜑}𝜓 ↔ ∃𝑦 ∈ 𝐴 (𝜑 ∧ 𝜒)) | ||
Theorem | abidnf 3408* | 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.) |
⊢ (Ⅎ𝑥𝐴 → {𝑧 ∣ ∀𝑥 𝑧 ∈ 𝐴} = 𝐴) | ||
Theorem | dedhb 3409* | A deduction theorem for converting the inference ⊢ Ⅎ𝑥𝐴 => ⊢ 𝜑 into a closed theorem. Use nfa1 2068 and nfab 2798 to eliminate the hypothesis of the substitution instance 𝜓 of the inference. For converting the inference form into a deduction form, abidnf 3408 is useful. (Contributed by NM, 8-Dec-2006.) |
⊢ (𝐴 = {𝑧 ∣ ∀𝑥 𝑧 ∈ 𝐴} → (𝜑 ↔ 𝜓)) & ⊢ 𝜓 ⇒ ⊢ (Ⅎ𝑥𝐴 → 𝜑) | ||
Theorem | eqeu 3410* | A condition which implies existential uniqueness. (Contributed by Jeff Hankins, 8-Sep-2009.) |
⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜓)) ⇒ ⊢ ((𝐴 ∈ 𝐵 ∧ 𝜓 ∧ ∀𝑥(𝜑 → 𝑥 = 𝐴)) → ∃!𝑥𝜑) | ||
Theorem | eueq 3411* | Equality has existential uniqueness. (Contributed by NM, 25-Nov-1994.) |
⊢ (𝐴 ∈ V ↔ ∃!𝑥 𝑥 = 𝐴) | ||
Theorem | eueq1 3412* | Equality has existential uniqueness. (Contributed by NM, 5-Apr-1995.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ∃!𝑥 𝑥 = 𝐴 | ||
Theorem | eueq2 3413* | Equality has existential uniqueness (split into 2 cases). (Contributed by NM, 5-Apr-1995.) |
⊢ 𝐴 ∈ V & ⊢ 𝐵 ∈ V ⇒ ⊢ ∃!𝑥((𝜑 ∧ 𝑥 = 𝐴) ∨ (¬ 𝜑 ∧ 𝑥 = 𝐵)) | ||
Theorem | eueq3 3414* | 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 & ⊢ ¬ (𝜑 ∧ 𝜓) ⇒ ⊢ ∃!𝑥((𝜑 ∧ 𝑥 = 𝐴) ∨ (¬ (𝜑 ∨ 𝜓) ∧ 𝑥 = 𝐵) ∨ (𝜓 ∧ 𝑥 = 𝐶)) | ||
Theorem | moeq 3415* | There is at most one set equal to a class. (Contributed by NM, 8-Mar-1995.) |
⊢ ∃*𝑥 𝑥 = 𝐴 | ||
Theorem | moeq3 3416* | "At most one" property of equality (split into 3 cases). (The first two hypotheses could be eliminated with longer proof.) (Contributed by NM, 23-Apr-1995.) |
⊢ 𝐵 ∈ V & ⊢ 𝐶 ∈ V & ⊢ ¬ (𝜑 ∧ 𝜓) ⇒ ⊢ ∃*𝑥((𝜑 ∧ 𝑥 = 𝐴) ∨ (¬ (𝜑 ∨ 𝜓) ∧ 𝑥 = 𝐵) ∨ (𝜓 ∧ 𝑥 = 𝐶)) | ||
Theorem | mosub 3417* | "At most one" remains true after substitution. (Contributed by NM, 9-Mar-1995.) |
⊢ ∃*𝑥𝜑 ⇒ ⊢ ∃*𝑥∃𝑦(𝑦 = 𝐴 ∧ 𝜑) | ||
Theorem | mo2icl 3418* | Theorem for inferring "at most one." (Contributed by NM, 17-Oct-1996.) |
⊢ (∀𝑥(𝜑 → 𝑥 = 𝐴) → ∃*𝑥𝜑) | ||
Theorem | mob2 3419* | Consequence of "at most one." (Contributed by NM, 2-Jan-2015.) |
⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜓)) ⇒ ⊢ ((𝐴 ∈ 𝐵 ∧ ∃*𝑥𝜑 ∧ 𝜑) → (𝑥 = 𝐴 ↔ 𝜓)) | ||
Theorem | moi2 3420* | Consequence of "at most one." (Contributed by NM, 29-Jun-2008.) |
⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (((𝐴 ∈ 𝐵 ∧ ∃*𝑥𝜑) ∧ (𝜑 ∧ 𝜓)) → 𝑥 = 𝐴) | ||
Theorem | mob 3421* | Equality implied by "at most one." (Contributed by NM, 18-Feb-2006.) |
⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝐵 → (𝜑 ↔ 𝜒)) ⇒ ⊢ (((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) ∧ ∃*𝑥𝜑 ∧ 𝜓) → (𝐴 = 𝐵 ↔ 𝜒)) | ||
Theorem | moi 3422* | Equality implied by "at most one." (Contributed by NM, 18-Feb-2006.) |
⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝐵 → (𝜑 ↔ 𝜒)) ⇒ ⊢ (((𝐴 ∈ 𝐶 ∧ 𝐵 ∈ 𝐷) ∧ ∃*𝑥𝜑 ∧ (𝜓 ∧ 𝜒)) → 𝐴 = 𝐵) | ||
Theorem | morex 3423* | Derive membership from uniqueness. (Contributed by Jeff Madsen, 2-Sep-2009.) |
⊢ 𝐵 ∈ V & ⊢ (𝑥 = 𝐵 → (𝜑 ↔ 𝜓)) ⇒ ⊢ ((∃𝑥 ∈ 𝐴 𝜑 ∧ ∃*𝑥𝜑) → (𝜓 → 𝐵 ∈ 𝐴)) | ||
Theorem | euxfr2 3424* | Transfer existential uniqueness from a variable 𝑥 to another variable 𝑦 contained in expression 𝐴. (Contributed by NM, 14-Nov-2004.) |
⊢ 𝐴 ∈ V & ⊢ ∃*𝑦 𝑥 = 𝐴 ⇒ ⊢ (∃!𝑥∃𝑦(𝑥 = 𝐴 ∧ 𝜑) ↔ ∃!𝑦𝜑) | ||
Theorem | euxfr 3425* | Transfer existential uniqueness from a variable 𝑥 to another variable 𝑦 contained in expression 𝐴. (Contributed by NM, 14-Nov-2004.) |
⊢ 𝐴 ∈ V & ⊢ ∃!𝑦 𝑥 = 𝐴 & ⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃!𝑥𝜑 ↔ ∃!𝑦𝜓) | ||
Theorem | euind 3426* | Existential uniqueness via an indirect equality. (Contributed by NM, 11-Oct-2010.) |
⊢ 𝐵 ∈ V & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝑦 → 𝐴 = 𝐵) ⇒ ⊢ ((∀𝑥∀𝑦((𝜑 ∧ 𝜓) → 𝐴 = 𝐵) ∧ ∃𝑥𝜑) → ∃!𝑧∀𝑥(𝜑 → 𝑧 = 𝐴)) | ||
Theorem | reu2 3427* | A way to express restricted uniqueness. (Contributed by NM, 22-Nov-1994.) |
⊢ (∃!𝑥 ∈ 𝐴 𝜑 ↔ (∃𝑥 ∈ 𝐴 𝜑 ∧ ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 ((𝜑 ∧ [𝑦 / 𝑥]𝜑) → 𝑥 = 𝑦))) | ||
Theorem | reu6 3428* | A way to express restricted uniqueness. (Contributed by NM, 20-Oct-2006.) |
⊢ (∃!𝑥 ∈ 𝐴 𝜑 ↔ ∃𝑦 ∈ 𝐴 ∀𝑥 ∈ 𝐴 (𝜑 ↔ 𝑥 = 𝑦)) | ||
Theorem | reu3 3429* | A way to express restricted uniqueness. (Contributed by NM, 24-Oct-2006.) |
⊢ (∃!𝑥 ∈ 𝐴 𝜑 ↔ (∃𝑥 ∈ 𝐴 𝜑 ∧ ∃𝑦 ∈ 𝐴 ∀𝑥 ∈ 𝐴 (𝜑 → 𝑥 = 𝑦))) | ||
Theorem | reu6i 3430* | A condition which implies existential uniqueness. (Contributed by Mario Carneiro, 2-Oct-2015.) |
⊢ ((𝐵 ∈ 𝐴 ∧ ∀𝑥 ∈ 𝐴 (𝜑 ↔ 𝑥 = 𝐵)) → ∃!𝑥 ∈ 𝐴 𝜑) | ||
Theorem | eqreu 3431* | A condition which implies existential uniqueness. (Contributed by Mario Carneiro, 2-Oct-2015.) |
⊢ (𝑥 = 𝐵 → (𝜑 ↔ 𝜓)) ⇒ ⊢ ((𝐵 ∈ 𝐴 ∧ 𝜓 ∧ ∀𝑥 ∈ 𝐴 (𝜑 → 𝑥 = 𝐵)) → ∃!𝑥 ∈ 𝐴 𝜑) | ||
Theorem | rmo4 3432* | Restricted "at most one" using implicit substitution. (Contributed by NM, 24-Oct-2006.) (Revised by NM, 16-Jun-2017.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃*𝑥 ∈ 𝐴 𝜑 ↔ ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 ((𝜑 ∧ 𝜓) → 𝑥 = 𝑦)) | ||
Theorem | reu4 3433* | Restricted uniqueness using implicit substitution. (Contributed by NM, 23-Nov-1994.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃!𝑥 ∈ 𝐴 𝜑 ↔ (∃𝑥 ∈ 𝐴 𝜑 ∧ ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 ((𝜑 ∧ 𝜓) → 𝑥 = 𝑦))) | ||
Theorem | reu7 3434* | Restricted uniqueness using implicit substitution. (Contributed by NM, 24-Oct-2006.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃!𝑥 ∈ 𝐴 𝜑 ↔ (∃𝑥 ∈ 𝐴 𝜑 ∧ ∃𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 (𝜓 → 𝑥 = 𝑦))) | ||
Theorem | reu8 3435* | Restricted uniqueness using implicit substitution. (Contributed by NM, 24-Oct-2006.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃!𝑥 ∈ 𝐴 𝜑 ↔ ∃𝑥 ∈ 𝐴 (𝜑 ∧ ∀𝑦 ∈ 𝐴 (𝜓 → 𝑥 = 𝑦))) | ||
Theorem | reu2eqd 3436* | Deduce equality from restricted uniqueness, deduction version. (Contributed by Thierry Arnoux, 27-Nov-2019.) |
⊢ (𝑥 = 𝐵 → (𝜓 ↔ 𝜒)) & ⊢ (𝑥 = 𝐶 → (𝜓 ↔ 𝜃)) & ⊢ (𝜑 → ∃!𝑥 ∈ 𝐴 𝜓) & ⊢ (𝜑 → 𝐵 ∈ 𝐴) & ⊢ (𝜑 → 𝐶 ∈ 𝐴) & ⊢ (𝜑 → 𝜒) & ⊢ (𝜑 → 𝜃) ⇒ ⊢ (𝜑 → 𝐵 = 𝐶) | ||
Theorem | reueq 3437* | Equality has existential uniqueness. (Contributed by Mario Carneiro, 1-Sep-2015.) |
⊢ (𝐵 ∈ 𝐴 ↔ ∃!𝑥 ∈ 𝐴 𝑥 = 𝐵) | ||
Theorem | rmoeq 3438* | Equality's restricted existential "at most one" property. (Contributed by Thierry Arnoux, 30-Mar-2018.) (Revised by AV, 27-Oct-2020.) (Proof shortened by NM, 29-Oct-2020.) |
⊢ ∃*𝑥 ∈ 𝐵 𝑥 = 𝐴 | ||
Theorem | rmoan 3439 | Restricted "at most one" still holds when a conjunct is added. (Contributed by NM, 16-Jun-2017.) |
⊢ (∃*𝑥 ∈ 𝐴 𝜑 → ∃*𝑥 ∈ 𝐴 (𝜓 ∧ 𝜑)) | ||
Theorem | rmoim 3440 | Restricted "at most one" is preserved through implication (note wff reversal). (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ (∀𝑥 ∈ 𝐴 (𝜑 → 𝜓) → (∃*𝑥 ∈ 𝐴 𝜓 → ∃*𝑥 ∈ 𝐴 𝜑)) | ||
Theorem | rmoimia 3441 | Restricted "at most one" is preserved through implication (note wff reversal). (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ (𝑥 ∈ 𝐴 → (𝜑 → 𝜓)) ⇒ ⊢ (∃*𝑥 ∈ 𝐴 𝜓 → ∃*𝑥 ∈ 𝐴 𝜑) | ||
Theorem | rmoimi2 3442 | Restricted "at most one" is preserved through implication (note wff reversal). (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ ∀𝑥((𝑥 ∈ 𝐴 ∧ 𝜑) → (𝑥 ∈ 𝐵 ∧ 𝜓)) ⇒ ⊢ (∃*𝑥 ∈ 𝐵 𝜓 → ∃*𝑥 ∈ 𝐴 𝜑) | ||
Theorem | 2reuswap 3443* | A condition allowing swap of uniqueness and existential quantifiers. (Contributed by Thierry Arnoux, 7-Apr-2017.) (Revised by NM, 16-Jun-2017.) |
⊢ (∀𝑥 ∈ 𝐴 ∃*𝑦 ∈ 𝐵 𝜑 → (∃!𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 𝜑 → ∃!𝑦 ∈ 𝐵 ∃𝑥 ∈ 𝐴 𝜑)) | ||
Theorem | reuind 3444* | Existential uniqueness via an indirect equality. (Contributed by NM, 16-Oct-2010.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝑦 → 𝐴 = 𝐵) ⇒ ⊢ ((∀𝑥∀𝑦(((𝐴 ∈ 𝐶 ∧ 𝜑) ∧ (𝐵 ∈ 𝐶 ∧ 𝜓)) → 𝐴 = 𝐵) ∧ ∃𝑥(𝐴 ∈ 𝐶 ∧ 𝜑)) → ∃!𝑧 ∈ 𝐶 ∀𝑥((𝐴 ∈ 𝐶 ∧ 𝜑) → 𝑧 = 𝐴)) | ||
Theorem | 2rmorex 3445* | Double restricted quantification with "at most one," analogous to 2moex 2572. (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ (∃*𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 𝜑 → ∀𝑦 ∈ 𝐵 ∃*𝑥 ∈ 𝐴 𝜑) | ||
Theorem | 2reu5lem1 3446* | Lemma for 2reu5 3449. Note that ∃!𝑥 ∈ 𝐴∃!𝑦 ∈ 𝐵𝜑 does not mean "there is exactly one 𝑥 in 𝐴 and exactly one 𝑦 in 𝐵 such that 𝜑 holds;" see comment for 2eu5 2586. (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ (∃!𝑥 ∈ 𝐴 ∃!𝑦 ∈ 𝐵 𝜑 ↔ ∃!𝑥∃!𝑦(𝑥 ∈ 𝐴 ∧ 𝑦 ∈ 𝐵 ∧ 𝜑)) | ||
Theorem | 2reu5lem2 3447* | Lemma for 2reu5 3449. (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ (∀𝑥 ∈ 𝐴 ∃*𝑦 ∈ 𝐵 𝜑 ↔ ∀𝑥∃*𝑦(𝑥 ∈ 𝐴 ∧ 𝑦 ∈ 𝐵 ∧ 𝜑)) | ||
Theorem | 2reu5lem3 3448* | Lemma for 2reu5 3449. 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 3559. (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ ((∃!𝑥 ∈ 𝐴 ∃!𝑦 ∈ 𝐵 𝜑 ∧ ∀𝑥 ∈ 𝐴 ∃*𝑦 ∈ 𝐵 𝜑) ↔ (∃𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 𝜑 ∧ ∃𝑧∃𝑤∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐵 (𝜑 → (𝑥 = 𝑧 ∧ 𝑦 = 𝑤)))) | ||
Theorem | 2reu5 3449* | Double restricted existential uniqueness in terms of restricted existential quantification and restricted universal quantification, analogous to 2eu5 2586 and reu3 3429. (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ ((∃!𝑥 ∈ 𝐴 ∃!𝑦 ∈ 𝐵 𝜑 ∧ ∀𝑥 ∈ 𝐴 ∃*𝑦 ∈ 𝐵 𝜑) ↔ (∃𝑥 ∈ 𝐴 ∃𝑦 ∈ 𝐵 𝜑 ∧ ∃𝑧 ∈ 𝐴 ∃𝑤 ∈ 𝐵 ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐵 (𝜑 → (𝑥 = 𝑧 ∧ 𝑦 = 𝑤)))) | ||
Theorem | nelrdva 3450* | Deduce negative membership from an implication. (Contributed by Thierry Arnoux, 27-Nov-2017.) |
⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → 𝑥 ≠ 𝐵) ⇒ ⊢ (𝜑 → ¬ 𝐵 ∈ 𝐴) | ||
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 ternary 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 3455 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 setvar 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 conditions of cdeqab1 3460 and cdeqab 3458. | ||
Syntax | wcdeq 3451 | 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(𝑥 = 𝑦 → 𝜑) | ||
Definition | df-cdeq 3452 | 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(𝑥 = 𝑦 → 𝜑) ↔ (𝑥 = 𝑦 → 𝜑)) | ||
Theorem | cdeqi 3453 | Deduce conditional equality. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ (𝑥 = 𝑦 → 𝜑) ⇒ ⊢ CondEq(𝑥 = 𝑦 → 𝜑) | ||
Theorem | cdeqri 3454 | Property of conditional equality. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ CondEq(𝑥 = 𝑦 → 𝜑) ⇒ ⊢ (𝑥 = 𝑦 → 𝜑) | ||
Theorem | cdeqth 3455 | Deduce conditional equality from a theorem. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ 𝜑 ⇒ ⊢ CondEq(𝑥 = 𝑦 → 𝜑) | ||
Theorem | cdeqnot 3456 | Distribute conditional equality over negation. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ CondEq(𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ CondEq(𝑥 = 𝑦 → (¬ 𝜑 ↔ ¬ 𝜓)) | ||
Theorem | cdeqal 3457* | Distribute conditional equality over quantification. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ CondEq(𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ CondEq(𝑥 = 𝑦 → (∀𝑧𝜑 ↔ ∀𝑧𝜓)) | ||
Theorem | cdeqab 3458* | Distribute conditional equality over abstraction. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ CondEq(𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ CondEq(𝑥 = 𝑦 → {𝑧 ∣ 𝜑} = {𝑧 ∣ 𝜓}) | ||
Theorem | cdeqal1 3459* | Distribute conditional equality over quantification. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ CondEq(𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ CondEq(𝑥 = 𝑦 → (∀𝑥𝜑 ↔ ∀𝑦𝜓)) | ||
Theorem | cdeqab1 3460* | Distribute conditional equality over abstraction. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ CondEq(𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ CondEq(𝑥 = 𝑦 → {𝑥 ∣ 𝜑} = {𝑦 ∣ 𝜓}) | ||
Theorem | cdeqim 3461 | Distribute conditional equality over implication. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ CondEq(𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) & ⊢ CondEq(𝑥 = 𝑦 → (𝜒 ↔ 𝜃)) ⇒ ⊢ CondEq(𝑥 = 𝑦 → ((𝜑 → 𝜒) ↔ (𝜓 → 𝜃))) | ||
Theorem | cdeqcv 3462 | Conditional equality for set-to-class promotion. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ CondEq(𝑥 = 𝑦 → 𝑥 = 𝑦) | ||
Theorem | cdeqeq 3463 | Distribute conditional equality over equality. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ CondEq(𝑥 = 𝑦 → 𝐴 = 𝐵) & ⊢ CondEq(𝑥 = 𝑦 → 𝐶 = 𝐷) ⇒ ⊢ CondEq(𝑥 = 𝑦 → (𝐴 = 𝐶 ↔ 𝐵 = 𝐷)) | ||
Theorem | cdeqel 3464 | Distribute conditional equality over elementhood. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ CondEq(𝑥 = 𝑦 → 𝐴 = 𝐵) & ⊢ CondEq(𝑥 = 𝑦 → 𝐶 = 𝐷) ⇒ ⊢ CondEq(𝑥 = 𝑦 → (𝐴 ∈ 𝐶 ↔ 𝐵 ∈ 𝐷)) | ||
Theorem | nfcdeq 3465* | 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(𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (𝜑 ↔ 𝜓) | ||
Theorem | nfccdeq 3466* | Variation of nfcdeq 3465 for classes. (Contributed by Mario Carneiro, 11-Aug-2016.) |
⊢ Ⅎ𝑥𝐴 & ⊢ CondEq(𝑥 = 𝑦 → 𝐴 = 𝐵) ⇒ ⊢ 𝐴 = 𝐵 | ||
Theorem | ru 3467 |
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 ssex 4835 asserting that 𝐴 is a set only when it is smaller than some other set 𝐵. However, Zermelo was then faced with a "chicken and egg" problem of how to show 𝐵 is a set, leading him to introduce the set-building axioms of Null Set 0ex 4823, Pairing prex 4939, Union uniex 6995, Power Set pwex 4878, and Infinity omex 8578 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 6014 (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 9335 and Cantor's Theorem canth 6648 are provably false! (See ncanth 6649 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 4814 replaces ax-rep 4804) with ax-sep 4814 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 our ZF set theory, every set is a member of the Russell class by elirrv 8542 (derived from the Axiom of Regularity), so for us the Russell class equals the universe V (theorem ruv 8545). See ruALT 8546 for an alternate proof of ru 3467 derived from that fact. (Contributed by NM, 7-Aug-1994.) |
⊢ {𝑥 ∣ 𝑥 ∉ 𝑥} ∉ V | ||
Syntax | wsbc 3468 | 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 [𝐴 / 𝑥]𝜑 | ||
Definition | df-sbc 3469 |
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 3495 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 3470 below). For example, if 𝐴 is a proper class, Quine's substitution of 𝐴 for 𝑦 in 0 ∈ 𝑦 evaluates to 0 ∈ 𝐴 rather than our falsehood. (This can be seen by substituting 𝐴, 𝑦, and 0 for alpha, beta, and gamma in Subcase 1 of Quine's discussion on p. 42.) Unfortunately, Quine's definition requires a recursive syntactic 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 3470, which holds for both our definition and Quine's, and from which we can derive a weaker version of df-sbc 3469 in the form of sbc8g 3476. 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 3469 and assert that [𝐴 / 𝑥]𝜑 is always false when 𝐴 is a proper class. The theorem sbc2or 3477 shows the apparently "strongest" statement we can make regarding behavior at proper classes if we start from dfsbcq 3470. The related definition df-csb 3567 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.) |
⊢ ([𝐴 / 𝑥]𝜑 ↔ 𝐴 ∈ {𝑥 ∣ 𝜑}) | ||
Theorem | dfsbcq 3470 |
Proper substitution of a class for a set in a wff given equal classes.
This is the essence of the sixth axiom of Frege, specifically Proposition
52 of [Frege1879] p. 50.
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 3469 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 3471 instead of df-sbc 3469. (dfsbcq2 3471 is needed because unlike Quine we do not overload the df-sb 1938 syntax.) As a consequence of these theorems, we can derive sbc8g 3476, which is a weaker version of df-sbc 3469 that leaves substitution undefined when 𝐴 is a proper class. However, it is often a nuisance to have to prove the sethood hypothesis of sbc8g 3476, so we will allow direct use of df-sbc 3469 after theorem sbc2or 3477 below. Proper substitution 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.) |
⊢ (𝐴 = 𝐵 → ([𝐴 / 𝑥]𝜑 ↔ [𝐵 / 𝑥]𝜑)) | ||
Theorem | dfsbcq2 3471 | 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 1938 and substitution for class variables df-sbc 3469. Unlike Quine, we use a different syntax for each in order to avoid overloading it. See remarks in dfsbcq 3470. (Contributed by NM, 31-Dec-2016.) |
⊢ (𝑦 = 𝐴 → ([𝑦 / 𝑥]𝜑 ↔ [𝐴 / 𝑥]𝜑)) | ||
Theorem | sbsbc 3472 | Show that df-sb 1938 and df-sbc 3469 are equivalent when the class term 𝐴 in df-sbc 3469 is a setvar variable. This theorem lets us reuse theorems based on df-sb 1938 for proofs involving df-sbc 3469. (Contributed by NM, 31-Dec-2016.) (Proof modification is discouraged.) |
⊢ ([𝑦 / 𝑥]𝜑 ↔ [𝑦 / 𝑥]𝜑) | ||
Theorem | sbceq1d 3473 | Equality theorem for class substitution. (Contributed by Mario Carneiro, 9-Feb-2017.) (Revised by NM, 30-Jun-2018.) |
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → ([𝐴 / 𝑥]𝜓 ↔ [𝐵 / 𝑥]𝜓)) | ||
Theorem | sbceq1dd 3474 | Equality theorem for class substitution. (Contributed by Mario Carneiro, 9-Feb-2017.) (Revised by NM, 30-Jun-2018.) |
⊢ (𝜑 → 𝐴 = 𝐵) & ⊢ (𝜑 → [𝐴 / 𝑥]𝜓) ⇒ ⊢ (𝜑 → [𝐵 / 𝑥]𝜓) | ||
Theorem | sbceqbid 3475* | Equality theorem for class substitution. (Contributed by Thierry Arnoux, 4-Sep-2018.) |
⊢ (𝜑 → 𝐴 = 𝐵) & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → ([𝐴 / 𝑥]𝜓 ↔ [𝐵 / 𝑥]𝜒)) | ||
Theorem | sbc8g 3476 | This is the closest we can get to df-sbc 3469 if we start from dfsbcq 3470 (see its comments) and dfsbcq2 3471. (Contributed by NM, 18-Nov-2008.) (Proof shortened by Andrew Salmon, 29-Jun-2011.) (Proof modification is discouraged.) |
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥]𝜑 ↔ 𝐴 ∈ {𝑥 ∣ 𝜑})) | ||
Theorem | sbc2or 3477* | 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 [𝐴 / 𝑥]𝜑 behavior at proper classes, matching the sbc5 3493 (false) and sbc6 3495 (true) conclusions. This is interesting since dfsbcq 3470 and dfsbcq2 3471 (from which it is derived) do not appear to say anything obvious about proper class behavior. Note that this theorem does not 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 𝑦 that 𝜑 or 𝐴 may contain. (Contributed by NM, 11-Oct-2004.) (Proof modification is discouraged.) |
⊢ (([𝐴 / 𝑥]𝜑 ↔ ∃𝑥(𝑥 = 𝐴 ∧ 𝜑)) ∨ ([𝐴 / 𝑥]𝜑 ↔ ∀𝑥(𝑥 = 𝐴 → 𝜑))) | ||
Theorem | sbcex 3478 | 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) | ||
Theorem | sbceq1a 3479 | Equality theorem for class substitution. Class version of sbequ12 2149. (Contributed by NM, 26-Sep-2003.) |
⊢ (𝑥 = 𝐴 → (𝜑 ↔ [𝐴 / 𝑥]𝜑)) | ||
Theorem | sbceq2a 3480 | Equality theorem for class substitution. Class version of sbequ12r 2150. (Contributed by NM, 4-Jan-2017.) |
⊢ (𝐴 = 𝑥 → ([𝐴 / 𝑥]𝜑 ↔ 𝜑)) | ||
Theorem | spsbc 3481 | 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. This is Frege's ninth axiom per Proposition 58 of [Frege1879] p. 51. See also stdpc4 2381 and rspsbc 3551. (Contributed by NM, 16-Jan-2004.) |
⊢ (𝐴 ∈ 𝑉 → (∀𝑥𝜑 → [𝐴 / 𝑥]𝜑)) | ||
Theorem | spsbcd 3482 | 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 2381 and rspsbc 3551. (Contributed by Mario Carneiro, 9-Feb-2017.) |
⊢ (𝜑 → 𝐴 ∈ 𝑉) & ⊢ (𝜑 → ∀𝑥𝜓) ⇒ ⊢ (𝜑 → [𝐴 / 𝑥]𝜓) | ||
Theorem | sbcth 3483 | A substitution into a theorem remains true (when 𝐴 is a set). (Contributed by NM, 5-Nov-2005.) |
⊢ 𝜑 ⇒ ⊢ (𝐴 ∈ 𝑉 → [𝐴 / 𝑥]𝜑) | ||
Theorem | sbcthdv 3484* | Deduction version of sbcth 3483. (Contributed by NM, 30-Nov-2005.) (Proof shortened by Andrew Salmon, 8-Jun-2011.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ ((𝜑 ∧ 𝐴 ∈ 𝑉) → [𝐴 / 𝑥]𝜓) | ||
Theorem | sbcid 3485 | An identity theorem for substitution. See sbid 2152. (Contributed by Mario Carneiro, 18-Feb-2017.) |
⊢ ([𝑥 / 𝑥]𝜑 ↔ 𝜑) | ||
Theorem | nfsbc1d 3486 | Deduction version of nfsbc1 3487. (Contributed by NM, 23-May-2006.) (Revised by Mario Carneiro, 12-Oct-2016.) |
⊢ (𝜑 → Ⅎ𝑥𝐴) ⇒ ⊢ (𝜑 → Ⅎ𝑥[𝐴 / 𝑥]𝜓) | ||
Theorem | nfsbc1 3487 | Bound-variable hypothesis builder for class substitution. (Contributed by NM, 5-Aug-1993.) (Revised by Mario Carneiro, 12-Oct-2016.) |
⊢ Ⅎ𝑥𝐴 ⇒ ⊢ Ⅎ𝑥[𝐴 / 𝑥]𝜑 | ||
Theorem | nfsbc1v 3488* | Bound-variable hypothesis builder for class substitution. (Contributed by Mario Carneiro, 12-Oct-2016.) |
⊢ Ⅎ𝑥[𝐴 / 𝑥]𝜑 | ||
Theorem | nfsbcd 3489 | Deduction version of nfsbc 3490. (Contributed by NM, 23-Nov-2005.) (Revised by Mario Carneiro, 12-Oct-2016.) |
⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝐴) & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → Ⅎ𝑥[𝐴 / 𝑦]𝜓) | ||
Theorem | nfsbc 3490 | Bound-variable hypothesis builder for class substitution. (Contributed by NM, 7-Sep-2014.) (Revised by Mario Carneiro, 12-Oct-2016.) |
⊢ Ⅎ𝑥𝐴 & ⊢ Ⅎ𝑥𝜑 ⇒ ⊢ Ⅎ𝑥[𝐴 / 𝑦]𝜑 | ||
Theorem | sbcco 3491* | A composition law for class substitution. (Contributed by NM, 26-Sep-2003.) (Revised by Mario Carneiro, 13-Oct-2016.) |
⊢ ([𝐴 / 𝑦][𝑦 / 𝑥]𝜑 ↔ [𝐴 / 𝑥]𝜑) | ||
Theorem | sbcco2 3492* | 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.) |
⊢ (𝑥 = 𝑦 → 𝐴 = 𝐵) ⇒ ⊢ ([𝑥 / 𝑦][𝐵 / 𝑥]𝜑 ↔ [𝐴 / 𝑥]𝜑) | ||
Theorem | sbc5 3493* | An equivalence for class substitution. (Contributed by NM, 23-Aug-1993.) (Revised by Mario Carneiro, 12-Oct-2016.) |
⊢ ([𝐴 / 𝑥]𝜑 ↔ ∃𝑥(𝑥 = 𝐴 ∧ 𝜑)) | ||
Theorem | sbc6g 3494* | An equivalence for class substitution. (Contributed by NM, 11-Oct-2004.) (Proof shortened by Andrew Salmon, 8-Jun-2011.) |
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥]𝜑 ↔ ∀𝑥(𝑥 = 𝐴 → 𝜑))) | ||
Theorem | sbc6 3495* | An equivalence for class substitution. (Contributed by NM, 23-Aug-1993.) (Proof shortened by Eric Schmidt, 17-Jan-2007.) |
⊢ 𝐴 ∈ V ⇒ ⊢ ([𝐴 / 𝑥]𝜑 ↔ ∀𝑥(𝑥 = 𝐴 → 𝜑)) | ||
Theorem | sbc7 3496* | An equivalence for class substitution in the spirit of df-clab 2638. Note that 𝑥 and 𝐴 don't have to be distinct. (Contributed by NM, 18-Nov-2008.) (Revised by Mario Carneiro, 13-Oct-2016.) |
⊢ ([𝐴 / 𝑥]𝜑 ↔ ∃𝑦(𝑦 = 𝐴 ∧ [𝑦 / 𝑥]𝜑)) | ||
Theorem | cbvsbc 3497 | Change bound variables in a wff substitution. (Contributed by Jeff Hankins, 19-Sep-2009.) (Proof shortened by Andrew Salmon, 8-Jun-2011.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ ([𝐴 / 𝑥]𝜑 ↔ [𝐴 / 𝑦]𝜓) | ||
Theorem | cbvsbcv 3498* | Change the bound variable of a class substitution using implicit substitution. (Contributed by NM, 30-Sep-2008.) (Revised by Mario Carneiro, 13-Oct-2016.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ ([𝐴 / 𝑥]𝜑 ↔ [𝐴 / 𝑦]𝜓) | ||
Theorem | sbciegft 3499* | Conversion of implicit substitution to explicit class substitution, using a bound-variable hypothesis instead of distinct variables. (Closed theorem version of sbciegf 3500.) (Contributed by NM, 10-Nov-2005.) (Revised by Mario Carneiro, 13-Oct-2016.) |
⊢ ((𝐴 ∈ 𝑉 ∧ Ⅎ𝑥𝜓 ∧ ∀𝑥(𝑥 = 𝐴 → (𝜑 ↔ 𝜓))) → ([𝐴 / 𝑥]𝜑 ↔ 𝜓)) | ||
Theorem | sbciegf 3500* | Conversion of implicit substitution to explicit class substitution. (Contributed by NM, 14-Dec-2005.) (Revised by Mario Carneiro, 13-Oct-2016.) |
⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝐴 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥]𝜑 ↔ 𝜓)) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |