Theorem List for Intuitionistic Logic Explorer - 3001-3100 *Has distinct variable
group(s)
Type | Label | Description |
Statement |
|
Theorem | sseldd 3001 |
Membership inference from subclass relationship. (Contributed by NM,
14-Dec-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐶 ∈ 𝐴) ⇒ ⊢ (𝜑 → 𝐶 ∈ 𝐵) |
|
Theorem | ssneld 3002 |
If a class is not in another class, it is also not in a subclass of that
class. Deduction form. (Contributed by David Moews, 1-May-2017.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵) ⇒ ⊢ (𝜑 → (¬ 𝐶 ∈ 𝐵 → ¬ 𝐶 ∈ 𝐴)) |
|
Theorem | ssneldd 3003 |
If an element is not in a class, it is also not in a subclass of that
class. Deduction form. (Contributed by David Moews, 1-May-2017.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → ¬ 𝐶 ∈ 𝐵) ⇒ ⊢ (𝜑 → ¬ 𝐶 ∈ 𝐴) |
|
Theorem | ssriv 3004* |
Inference rule based on subclass definition. (Contributed by NM,
5-Aug-1993.)
|
⊢ (𝑥 ∈ 𝐴 → 𝑥 ∈ 𝐵) ⇒ ⊢ 𝐴 ⊆ 𝐵 |
|
Theorem | ssrd 3005 |
Deduction rule based on subclass definition. (Contributed by Thierry
Arnoux, 8-Mar-2017.)
|
⊢ Ⅎ𝑥𝜑
& ⊢ Ⅎ𝑥𝐴
& ⊢ Ⅎ𝑥𝐵
& ⊢ (𝜑 → (𝑥 ∈ 𝐴 → 𝑥 ∈ 𝐵)) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐵) |
|
Theorem | ssrdv 3006* |
Deduction rule based on subclass definition. (Contributed by NM,
15-Nov-1995.)
|
⊢ (𝜑 → (𝑥 ∈ 𝐴 → 𝑥 ∈ 𝐵)) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐵) |
|
Theorem | sstr2 3007 |
Transitivity of subclasses. Exercise 5 of [TakeutiZaring] p. 17.
(Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew Salmon,
14-Jun-2011.)
|
⊢ (𝐴 ⊆ 𝐵 → (𝐵 ⊆ 𝐶 → 𝐴 ⊆ 𝐶)) |
|
Theorem | sstr 3008 |
Transitivity of subclasses. Theorem 6 of [Suppes] p. 23. (Contributed by
NM, 5-Sep-2003.)
|
⊢ ((𝐴 ⊆ 𝐵 ∧ 𝐵 ⊆ 𝐶) → 𝐴 ⊆ 𝐶) |
|
Theorem | sstri 3009 |
Subclass transitivity inference. (Contributed by NM, 5-May-2000.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ 𝐵 ⊆ 𝐶 ⇒ ⊢ 𝐴 ⊆ 𝐶 |
|
Theorem | sstrd 3010 |
Subclass transitivity deduction. (Contributed by NM, 2-Jun-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl5ss 3011 |
Subclass transitivity deduction. (Contributed by NM, 6-Feb-2014.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl6ss 3012 |
Subclass transitivity deduction. (Contributed by Jonathan Ben-Naim,
3-Jun-2011.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ 𝐵 ⊆ 𝐶 ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | sylan9ss 3013 |
A subclass transitivity deduction. (Contributed by NM, 27-Sep-2004.)
(Proof shortened by Andrew Salmon, 14-Jun-2011.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜓 → 𝐵 ⊆ 𝐶) ⇒ ⊢ ((𝜑 ∧ 𝜓) → 𝐴 ⊆ 𝐶) |
|
Theorem | sylan9ssr 3014 |
A subclass transitivity deduction. (Contributed by NM, 27-Sep-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜓 → 𝐵 ⊆ 𝐶) ⇒ ⊢ ((𝜓 ∧ 𝜑) → 𝐴 ⊆ 𝐶) |
|
Theorem | eqss 3015 |
The subclass relationship is antisymmetric. Compare Theorem 4 of
[Suppes] p. 22. (Contributed by NM,
5-Aug-1993.)
|
⊢ (𝐴 = 𝐵 ↔ (𝐴 ⊆ 𝐵 ∧ 𝐵 ⊆ 𝐴)) |
|
Theorem | eqssi 3016 |
Infer equality from two subclass relationships. Compare Theorem 4 of
[Suppes] p. 22. (Contributed by NM,
9-Sep-1993.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ 𝐵 ⊆ 𝐴 ⇒ ⊢ 𝐴 = 𝐵 |
|
Theorem | eqssd 3017 |
Equality deduction from two subclass relationships. Compare Theorem 4
of [Suppes] p. 22. (Contributed by NM,
27-Jun-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐵 ⊆ 𝐴) ⇒ ⊢ (𝜑 → 𝐴 = 𝐵) |
|
Theorem | eqrd 3018 |
Deduce equality of classes from equivalence of membership. (Contributed
by Thierry Arnoux, 21-Mar-2017.)
|
⊢ Ⅎ𝑥𝜑
& ⊢ Ⅎ𝑥𝐴
& ⊢ Ⅎ𝑥𝐵
& ⊢ (𝜑 → (𝑥 ∈ 𝐴 ↔ 𝑥 ∈ 𝐵)) ⇒ ⊢ (𝜑 → 𝐴 = 𝐵) |
|
Theorem | ssid 3019 |
Any class is a subclass of itself. Exercise 10 of [TakeutiZaring]
p. 18. (Contributed by NM, 5-Aug-1993.) (Proof shortened by Andrew
Salmon, 14-Jun-2011.)
|
⊢ 𝐴 ⊆ 𝐴 |
|
Theorem | ssv 3020 |
Any class is a subclass of the universal class. (Contributed by NM,
31-Oct-1995.)
|
⊢ 𝐴 ⊆ V |
|
Theorem | sseq1 3021 |
Equality theorem for subclasses. (Contributed by NM, 5-Aug-1993.) (Proof
shortened by Andrew Salmon, 21-Jun-2011.)
|
⊢ (𝐴 = 𝐵 → (𝐴 ⊆ 𝐶 ↔ 𝐵 ⊆ 𝐶)) |
|
Theorem | sseq2 3022 |
Equality theorem for the subclass relationship. (Contributed by NM,
25-Jun-1998.)
|
⊢ (𝐴 = 𝐵 → (𝐶 ⊆ 𝐴 ↔ 𝐶 ⊆ 𝐵)) |
|
Theorem | sseq12 3023 |
Equality theorem for the subclass relationship. (Contributed by NM,
31-May-1999.)
|
⊢ ((𝐴 = 𝐵 ∧ 𝐶 = 𝐷) → (𝐴 ⊆ 𝐶 ↔ 𝐵 ⊆ 𝐷)) |
|
Theorem | sseq1i 3024 |
An equality inference for the subclass relationship. (Contributed by
NM, 18-Aug-1993.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ (𝐴 ⊆ 𝐶 ↔ 𝐵 ⊆ 𝐶) |
|
Theorem | sseq2i 3025 |
An equality inference for the subclass relationship. (Contributed by
NM, 30-Aug-1993.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ (𝐶 ⊆ 𝐴 ↔ 𝐶 ⊆ 𝐵) |
|
Theorem | sseq12i 3026 |
An equality inference for the subclass relationship. (Contributed by
NM, 31-May-1999.) (Proof shortened by Eric Schmidt, 26-Jan-2007.)
|
⊢ 𝐴 = 𝐵
& ⊢ 𝐶 = 𝐷 ⇒ ⊢ (𝐴 ⊆ 𝐶 ↔ 𝐵 ⊆ 𝐷) |
|
Theorem | sseq1d 3027 |
An equality deduction for the subclass relationship. (Contributed by
NM, 14-Aug-1994.)
|
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐴 ⊆ 𝐶 ↔ 𝐵 ⊆ 𝐶)) |
|
Theorem | sseq2d 3028 |
An equality deduction for the subclass relationship. (Contributed by
NM, 14-Aug-1994.)
|
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐶 ⊆ 𝐴 ↔ 𝐶 ⊆ 𝐵)) |
|
Theorem | sseq12d 3029 |
An equality deduction for the subclass relationship. (Contributed by
NM, 31-May-1999.)
|
⊢ (𝜑 → 𝐴 = 𝐵)
& ⊢ (𝜑 → 𝐶 = 𝐷) ⇒ ⊢ (𝜑 → (𝐴 ⊆ 𝐶 ↔ 𝐵 ⊆ 𝐷)) |
|
Theorem | eqsstri 3030 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 16-Jul-1995.)
|
⊢ 𝐴 = 𝐵
& ⊢ 𝐵 ⊆ 𝐶 ⇒ ⊢ 𝐴 ⊆ 𝐶 |
|
Theorem | eqsstr3i 3031 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 19-Oct-1999.)
|
⊢ 𝐵 = 𝐴
& ⊢ 𝐵 ⊆ 𝐶 ⇒ ⊢ 𝐴 ⊆ 𝐶 |
|
Theorem | sseqtri 3032 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 28-Jul-1995.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ 𝐵 = 𝐶 ⇒ ⊢ 𝐴 ⊆ 𝐶 |
|
Theorem | sseqtr4i 3033 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 4-Apr-1995.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ 𝐶 = 𝐵 ⇒ ⊢ 𝐴 ⊆ 𝐶 |
|
Theorem | eqsstrd 3034 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 25-Apr-2004.)
|
⊢ (𝜑 → 𝐴 = 𝐵)
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | eqsstr3d 3035 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 25-Apr-2004.)
|
⊢ (𝜑 → 𝐵 = 𝐴)
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | sseqtrd 3036 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 25-Apr-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐵 = 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | sseqtr4d 3037 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 25-Apr-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐶 = 𝐵) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | 3sstr3i 3038 |
Substitution of equality in both sides of a subclass relationship.
(Contributed by NM, 13-Jan-1996.) (Proof shortened by Eric Schmidt,
26-Jan-2007.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ 𝐴 = 𝐶
& ⊢ 𝐵 = 𝐷 ⇒ ⊢ 𝐶 ⊆ 𝐷 |
|
Theorem | 3sstr4i 3039 |
Substitution of equality in both sides of a subclass relationship.
(Contributed by NM, 13-Jan-1996.) (Proof shortened by Eric Schmidt,
26-Jan-2007.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ 𝐶 = 𝐴
& ⊢ 𝐷 = 𝐵 ⇒ ⊢ 𝐶 ⊆ 𝐷 |
|
Theorem | 3sstr3g 3040 |
Substitution of equality into both sides of a subclass relationship.
(Contributed by NM, 1-Oct-2000.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ 𝐴 = 𝐶
& ⊢ 𝐵 = 𝐷 ⇒ ⊢ (𝜑 → 𝐶 ⊆ 𝐷) |
|
Theorem | 3sstr4g 3041 |
Substitution of equality into both sides of a subclass relationship.
(Contributed by NM, 16-Aug-1994.) (Proof shortened by Eric Schmidt,
26-Jan-2007.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ 𝐶 = 𝐴
& ⊢ 𝐷 = 𝐵 ⇒ ⊢ (𝜑 → 𝐶 ⊆ 𝐷) |
|
Theorem | 3sstr3d 3042 |
Substitution of equality into both sides of a subclass relationship.
(Contributed by NM, 1-Oct-2000.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐴 = 𝐶)
& ⊢ (𝜑 → 𝐵 = 𝐷) ⇒ ⊢ (𝜑 → 𝐶 ⊆ 𝐷) |
|
Theorem | 3sstr4d 3043 |
Substitution of equality into both sides of a subclass relationship.
(Contributed by NM, 30-Nov-1995.) (Proof shortened by Eric Schmidt,
26-Jan-2007.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐶 = 𝐴)
& ⊢ (𝜑 → 𝐷 = 𝐵) ⇒ ⊢ (𝜑 → 𝐶 ⊆ 𝐷) |
|
Theorem | syl5eqss 3044 |
B chained subclass and equality deduction. (Contributed by NM,
25-Apr-2004.)
|
⊢ 𝐴 = 𝐵
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl5eqssr 3045 |
B chained subclass and equality deduction. (Contributed by NM,
25-Apr-2004.)
|
⊢ 𝐵 = 𝐴
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl6sseq 3046 |
A chained subclass and equality deduction. (Contributed by NM,
25-Apr-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ 𝐵 = 𝐶 ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl6sseqr 3047 |
A chained subclass and equality deduction. (Contributed by NM,
25-Apr-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ 𝐶 = 𝐵 ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl5sseq 3048 |
Subclass transitivity deduction. (Contributed by Jonathan Ben-Naim,
3-Jun-2011.)
|
⊢ 𝐵 ⊆ 𝐴
& ⊢ (𝜑 → 𝐴 = 𝐶) ⇒ ⊢ (𝜑 → 𝐵 ⊆ 𝐶) |
|
Theorem | syl5sseqr 3049 |
Subclass transitivity deduction. (Contributed by Jonathan Ben-Naim,
3-Jun-2011.)
|
⊢ 𝐵 ⊆ 𝐴
& ⊢ (𝜑 → 𝐶 = 𝐴) ⇒ ⊢ (𝜑 → 𝐵 ⊆ 𝐶) |
|
Theorem | syl6eqss 3050 |
A chained subclass and equality deduction. (Contributed by Mario
Carneiro, 2-Jan-2017.)
|
⊢ (𝜑 → 𝐴 = 𝐵)
& ⊢ 𝐵 ⊆ 𝐶 ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl6eqssr 3051 |
A chained subclass and equality deduction. (Contributed by Mario
Carneiro, 2-Jan-2017.)
|
⊢ (𝜑 → 𝐵 = 𝐴)
& ⊢ 𝐵 ⊆ 𝐶 ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | eqimss 3052 |
Equality implies the subclass relation. (Contributed by NM, 5-Aug-1993.)
(Proof shortened by Andrew Salmon, 21-Jun-2011.)
|
⊢ (𝐴 = 𝐵 → 𝐴 ⊆ 𝐵) |
|
Theorem | eqimss2 3053 |
Equality implies the subclass relation. (Contributed by NM,
23-Nov-2003.)
|
⊢ (𝐵 = 𝐴 → 𝐴 ⊆ 𝐵) |
|
Theorem | eqimssi 3054 |
Infer subclass relationship from equality. (Contributed by NM,
6-Jan-2007.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ 𝐴 ⊆ 𝐵 |
|
Theorem | eqimss2i 3055 |
Infer subclass relationship from equality. (Contributed by NM,
7-Jan-2007.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ 𝐵 ⊆ 𝐴 |
|
Theorem | nssne1 3056 |
Two classes are different if they don't include the same class.
(Contributed by NM, 23-Apr-2015.)
|
⊢ ((𝐴 ⊆ 𝐵 ∧ ¬ 𝐴 ⊆ 𝐶) → 𝐵 ≠ 𝐶) |
|
Theorem | nssne2 3057 |
Two classes are different if they are not subclasses of the same class.
(Contributed by NM, 23-Apr-2015.)
|
⊢ ((𝐴 ⊆ 𝐶 ∧ ¬ 𝐵 ⊆ 𝐶) → 𝐴 ≠ 𝐵) |
|
Theorem | nssr 3058* |
Negation of subclass relationship. One direction of Exercise 13 of
[TakeutiZaring] p. 18.
(Contributed by Jim Kingdon, 15-Jul-2018.)
|
⊢ (∃𝑥(𝑥 ∈ 𝐴 ∧ ¬ 𝑥 ∈ 𝐵) → ¬ 𝐴 ⊆ 𝐵) |
|
Theorem | ssralv 3059* |
Quantification restricted to a subclass. (Contributed by NM,
11-Mar-2006.)
|
⊢ (𝐴 ⊆ 𝐵 → (∀𝑥 ∈ 𝐵 𝜑 → ∀𝑥 ∈ 𝐴 𝜑)) |
|
Theorem | ssrexv 3060* |
Existential quantification restricted to a subclass. (Contributed by
NM, 11-Jan-2007.)
|
⊢ (𝐴 ⊆ 𝐵 → (∃𝑥 ∈ 𝐴 𝜑 → ∃𝑥 ∈ 𝐵 𝜑)) |
|
Theorem | ralss 3061* |
Restricted universal quantification on a subset in terms of superset.
(Contributed by Stefan O'Rear, 3-Apr-2015.)
|
⊢ (𝐴 ⊆ 𝐵 → (∀𝑥 ∈ 𝐴 𝜑 ↔ ∀𝑥 ∈ 𝐵 (𝑥 ∈ 𝐴 → 𝜑))) |
|
Theorem | rexss 3062* |
Restricted existential quantification on a subset in terms of superset.
(Contributed by Stefan O'Rear, 3-Apr-2015.)
|
⊢ (𝐴 ⊆ 𝐵 → (∃𝑥 ∈ 𝐴 𝜑 ↔ ∃𝑥 ∈ 𝐵 (𝑥 ∈ 𝐴 ∧ 𝜑))) |
|
Theorem | ss2ab 3063 |
Class abstractions in a subclass relationship. (Contributed by NM,
3-Jul-1994.)
|
⊢ ({𝑥 ∣ 𝜑} ⊆ {𝑥 ∣ 𝜓} ↔ ∀𝑥(𝜑 → 𝜓)) |
|
Theorem | abss 3064* |
Class abstraction in a subclass relationship. (Contributed by NM,
16-Aug-2006.)
|
⊢ ({𝑥 ∣ 𝜑} ⊆ 𝐴 ↔ ∀𝑥(𝜑 → 𝑥 ∈ 𝐴)) |
|
Theorem | ssab 3065* |
Subclass of a class abstraction. (Contributed by NM, 16-Aug-2006.)
|
⊢ (𝐴 ⊆ {𝑥 ∣ 𝜑} ↔ ∀𝑥(𝑥 ∈ 𝐴 → 𝜑)) |
|
Theorem | ssabral 3066* |
The relation for a subclass of a class abstraction is equivalent to
restricted quantification. (Contributed by NM, 6-Sep-2006.)
|
⊢ (𝐴 ⊆ {𝑥 ∣ 𝜑} ↔ ∀𝑥 ∈ 𝐴 𝜑) |
|
Theorem | ss2abi 3067 |
Inference of abstraction subclass from implication. (Contributed by NM,
31-Mar-1995.)
|
⊢ (𝜑 → 𝜓) ⇒ ⊢ {𝑥 ∣ 𝜑} ⊆ {𝑥 ∣ 𝜓} |
|
Theorem | ss2abdv 3068* |
Deduction of abstraction subclass from implication. (Contributed by NM,
29-Jul-2011.)
|
⊢ (𝜑 → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → {𝑥 ∣ 𝜓} ⊆ {𝑥 ∣ 𝜒}) |
|
Theorem | abssdv 3069* |
Deduction of abstraction subclass from implication. (Contributed by NM,
20-Jan-2006.)
|
⊢ (𝜑 → (𝜓 → 𝑥 ∈ 𝐴)) ⇒ ⊢ (𝜑 → {𝑥 ∣ 𝜓} ⊆ 𝐴) |
|
Theorem | abssi 3070* |
Inference of abstraction subclass from implication. (Contributed by NM,
20-Jan-2006.)
|
⊢ (𝜑 → 𝑥 ∈ 𝐴) ⇒ ⊢ {𝑥 ∣ 𝜑} ⊆ 𝐴 |
|
Theorem | ss2rab 3071 |
Restricted abstraction classes in a subclass relationship. (Contributed
by NM, 30-May-1999.)
|
⊢ ({𝑥 ∈ 𝐴 ∣ 𝜑} ⊆ {𝑥 ∈ 𝐴 ∣ 𝜓} ↔ ∀𝑥 ∈ 𝐴 (𝜑 → 𝜓)) |
|
Theorem | rabss 3072* |
Restricted class abstraction in a subclass relationship. (Contributed
by NM, 16-Aug-2006.)
|
⊢ ({𝑥 ∈ 𝐴 ∣ 𝜑} ⊆ 𝐵 ↔ ∀𝑥 ∈ 𝐴 (𝜑 → 𝑥 ∈ 𝐵)) |
|
Theorem | ssrab 3073* |
Subclass of a restricted class abstraction. (Contributed by NM,
16-Aug-2006.)
|
⊢ (𝐵 ⊆ {𝑥 ∈ 𝐴 ∣ 𝜑} ↔ (𝐵 ⊆ 𝐴 ∧ ∀𝑥 ∈ 𝐵 𝜑)) |
|
Theorem | ssrabdv 3074* |
Subclass of a restricted class abstraction (deduction rule).
(Contributed by NM, 31-Aug-2006.)
|
⊢ (𝜑 → 𝐵 ⊆ 𝐴)
& ⊢ ((𝜑 ∧ 𝑥 ∈ 𝐵) → 𝜓) ⇒ ⊢ (𝜑 → 𝐵 ⊆ {𝑥 ∈ 𝐴 ∣ 𝜓}) |
|
Theorem | rabssdv 3075* |
Subclass of a restricted class abstraction (deduction rule).
(Contributed by NM, 2-Feb-2015.)
|
⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴 ∧ 𝜓) → 𝑥 ∈ 𝐵) ⇒ ⊢ (𝜑 → {𝑥 ∈ 𝐴 ∣ 𝜓} ⊆ 𝐵) |
|
Theorem | ss2rabdv 3076* |
Deduction of restricted abstraction subclass from implication.
(Contributed by NM, 30-May-2006.)
|
⊢ ((𝜑 ∧ 𝑥 ∈ 𝐴) → (𝜓 → 𝜒)) ⇒ ⊢ (𝜑 → {𝑥 ∈ 𝐴 ∣ 𝜓} ⊆ {𝑥 ∈ 𝐴 ∣ 𝜒}) |
|
Theorem | ss2rabi 3077 |
Inference of restricted abstraction subclass from implication.
(Contributed by NM, 14-Oct-1999.)
|
⊢ (𝑥 ∈ 𝐴 → (𝜑 → 𝜓)) ⇒ ⊢ {𝑥 ∈ 𝐴 ∣ 𝜑} ⊆ {𝑥 ∈ 𝐴 ∣ 𝜓} |
|
Theorem | rabss2 3078* |
Subclass law for restricted abstraction. (Contributed by NM,
18-Dec-2004.) (Proof shortened by Andrew Salmon, 26-Jun-2011.)
|
⊢ (𝐴 ⊆ 𝐵 → {𝑥 ∈ 𝐴 ∣ 𝜑} ⊆ {𝑥 ∈ 𝐵 ∣ 𝜑}) |
|
Theorem | ssab2 3079* |
Subclass relation for the restriction of a class abstraction.
(Contributed by NM, 31-Mar-1995.)
|
⊢ {𝑥 ∣ (𝑥 ∈ 𝐴 ∧ 𝜑)} ⊆ 𝐴 |
|
Theorem | ssrab2 3080* |
Subclass relation for a restricted class. (Contributed by NM,
19-Mar-1997.)
|
⊢ {𝑥 ∈ 𝐴 ∣ 𝜑} ⊆ 𝐴 |
|
Theorem | ssrabeq 3081* |
If the restricting class of a restricted class abstraction is a subset
of this restricted class abstraction, it is equal to this restricted
class abstraction. (Contributed by Alexander van der Vekens,
31-Dec-2017.)
|
⊢ (𝑉 ⊆ {𝑥 ∈ 𝑉 ∣ 𝜑} ↔ 𝑉 = {𝑥 ∈ 𝑉 ∣ 𝜑}) |
|
Theorem | rabssab 3082 |
A restricted class is a subclass of the corresponding unrestricted class.
(Contributed by Mario Carneiro, 23-Dec-2016.)
|
⊢ {𝑥 ∈ 𝐴 ∣ 𝜑} ⊆ {𝑥 ∣ 𝜑} |
|
Theorem | uniiunlem 3083* |
A subset relationship useful for converting union to indexed union using
dfiun2 or dfiun2g and intersection to indexed intersection using
dfiin2 . (Contributed by NM, 5-Oct-2006.) (Proof shortened by Mario
Carneiro, 26-Sep-2015.)
|
⊢ (∀𝑥 ∈ 𝐴 𝐵 ∈ 𝐷 → (∀𝑥 ∈ 𝐴 𝐵 ∈ 𝐶 ↔ {𝑦 ∣ ∃𝑥 ∈ 𝐴 𝑦 = 𝐵} ⊆ 𝐶)) |
|
2.1.13 The difference, union, and intersection
of two classes
|
|
2.1.13.1 The difference of two
classes
|
|
Theorem | difeq1 3084 |
Equality theorem for class difference. (Contributed by NM,
10-Feb-1997.) (Proof shortened by Andrew Salmon, 26-Jun-2011.)
|
⊢ (𝐴 = 𝐵 → (𝐴 ∖ 𝐶) = (𝐵 ∖ 𝐶)) |
|
Theorem | difeq2 3085 |
Equality theorem for class difference. (Contributed by NM,
10-Feb-1997.) (Proof shortened by Andrew Salmon, 26-Jun-2011.)
|
⊢ (𝐴 = 𝐵 → (𝐶 ∖ 𝐴) = (𝐶 ∖ 𝐵)) |
|
Theorem | difeq12 3086 |
Equality theorem for class difference. (Contributed by FL,
31-Aug-2009.)
|
⊢ ((𝐴 = 𝐵 ∧ 𝐶 = 𝐷) → (𝐴 ∖ 𝐶) = (𝐵 ∖ 𝐷)) |
|
Theorem | difeq1i 3087 |
Inference adding difference to the right in a class equality.
(Contributed by NM, 15-Nov-2002.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ (𝐴 ∖ 𝐶) = (𝐵 ∖ 𝐶) |
|
Theorem | difeq2i 3088 |
Inference adding difference to the left in a class equality.
(Contributed by NM, 15-Nov-2002.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ (𝐶 ∖ 𝐴) = (𝐶 ∖ 𝐵) |
|
Theorem | difeq12i 3089 |
Equality inference for class difference. (Contributed by NM,
29-Aug-2004.)
|
⊢ 𝐴 = 𝐵
& ⊢ 𝐶 = 𝐷 ⇒ ⊢ (𝐴 ∖ 𝐶) = (𝐵 ∖ 𝐷) |
|
Theorem | difeq1d 3090 |
Deduction adding difference to the right in a class equality.
(Contributed by NM, 15-Nov-2002.)
|
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐴 ∖ 𝐶) = (𝐵 ∖ 𝐶)) |
|
Theorem | difeq2d 3091 |
Deduction adding difference to the left in a class equality.
(Contributed by NM, 15-Nov-2002.)
|
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐶 ∖ 𝐴) = (𝐶 ∖ 𝐵)) |
|
Theorem | difeq12d 3092 |
Equality deduction for class difference. (Contributed by FL,
29-May-2014.)
|
⊢ (𝜑 → 𝐴 = 𝐵)
& ⊢ (𝜑 → 𝐶 = 𝐷) ⇒ ⊢ (𝜑 → (𝐴 ∖ 𝐶) = (𝐵 ∖ 𝐷)) |
|
Theorem | difeqri 3093* |
Inference from membership to difference. (Contributed by NM,
17-May-1998.) (Proof shortened by Andrew Salmon, 26-Jun-2011.)
|
⊢ ((𝑥 ∈ 𝐴 ∧ ¬ 𝑥 ∈ 𝐵) ↔ 𝑥 ∈ 𝐶) ⇒ ⊢ (𝐴 ∖ 𝐵) = 𝐶 |
|
Theorem | nfdif 3094 |
Bound-variable hypothesis builder for class difference. (Contributed by
NM, 3-Dec-2003.) (Revised by Mario Carneiro, 13-Oct-2016.)
|
⊢ Ⅎ𝑥𝐴
& ⊢ Ⅎ𝑥𝐵 ⇒ ⊢ Ⅎ𝑥(𝐴 ∖ 𝐵) |
|
Theorem | eldifi 3095 |
Implication of membership in a class difference. (Contributed by NM,
29-Apr-1994.)
|
⊢ (𝐴 ∈ (𝐵 ∖ 𝐶) → 𝐴 ∈ 𝐵) |
|
Theorem | eldifn 3096 |
Implication of membership in a class difference. (Contributed by NM,
3-May-1994.)
|
⊢ (𝐴 ∈ (𝐵 ∖ 𝐶) → ¬ 𝐴 ∈ 𝐶) |
|
Theorem | elndif 3097 |
A set does not belong to a class excluding it. (Contributed by NM,
27-Jun-1994.)
|
⊢ (𝐴 ∈ 𝐵 → ¬ 𝐴 ∈ (𝐶 ∖ 𝐵)) |
|
Theorem | difdif 3098 |
Double class difference. Exercise 11 of [TakeutiZaring] p. 22.
(Contributed by NM, 17-May-1998.)
|
⊢ (𝐴 ∖ (𝐵 ∖ 𝐴)) = 𝐴 |
|
Theorem | difss 3099 |
Subclass relationship for class difference. Exercise 14 of
[TakeutiZaring] p. 22.
(Contributed by NM, 29-Apr-1994.)
|
⊢ (𝐴 ∖ 𝐵) ⊆ 𝐴 |
|
Theorem | difssd 3100 |
A difference of two classes is contained in the minuend. Deduction form
of difss 3099. (Contributed by David Moews, 1-May-2017.)
|
⊢ (𝜑 → (𝐴 ∖ 𝐵) ⊆ 𝐴) |