Theorem List for Intuitionistic Logic Explorer - 2901-3000 *Has distinct variable
group(s)
Type | Label | Description |
Statement |
|
Theorem | csbnest1g 2901 |
Nest the composition of two substitutions. (Contributed by NM,
23-May-2006.) (Proof shortened by Mario Carneiro, 11-Nov-2016.)
|
⊢ (𝐴 ∈ 𝑉 → ⦋𝐴 / 𝑥⦌⦋𝐵 / 𝑥⦌𝐶 = ⦋⦋𝐴 / 𝑥⦌𝐵 / 𝑥⦌𝐶) |
|
Theorem | csbidmg 2902* |
Idempotent law for class substitutions. (Contributed by NM,
1-Mar-2008.)
|
⊢ (𝐴 ∈ 𝑉 → ⦋𝐴 / 𝑥⦌⦋𝐴 / 𝑥⦌𝐵 = ⦋𝐴 / 𝑥⦌𝐵) |
|
Theorem | sbcco3g 2903* |
Composition of two substitutions. (Contributed by NM, 27-Nov-2005.)
(Revised by Mario Carneiro, 11-Nov-2016.)
|
⊢ (𝑥 = 𝐴 → 𝐵 = 𝐶) ⇒ ⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥][𝐵 / 𝑦]𝜑 ↔ [𝐶 / 𝑦]𝜑)) |
|
Theorem | csbco3g 2904* |
Composition of two class substitutions. (Contributed by NM,
27-Nov-2005.) (Revised by Mario Carneiro, 11-Nov-2016.)
|
⊢ (𝑥 = 𝐴 → 𝐵 = 𝐶) ⇒ ⊢ (𝐴 ∈ 𝑉 → ⦋𝐴 / 𝑥⦌⦋𝐵 / 𝑦⦌𝐷 = ⦋𝐶 / 𝑦⦌𝐷) |
|
Theorem | rspcsbela 2905* |
Special case related to rspsbc 2840. (Contributed by NM, 10-Dec-2005.)
(Proof shortened by Eric Schmidt, 17-Jan-2007.)
|
⊢ ((𝐴 ∈ 𝐵 ∧ ∀𝑥 ∈ 𝐵 𝐶 ∈ 𝐷) → ⦋𝐴 / 𝑥⦌𝐶 ∈ 𝐷) |
|
Theorem | sbnfc2 2906* |
Two ways of expressing "𝑥 is (effectively) not free in 𝐴."
(Contributed by Mario Carneiro, 14-Oct-2016.)
|
⊢ (Ⅎ𝑥𝐴 ↔ ∀𝑦∀𝑧⦋𝑦 / 𝑥⦌𝐴 = ⦋𝑧 / 𝑥⦌𝐴) |
|
Theorem | csbabg 2907* |
Move substitution into a class abstraction. (Contributed by NM,
13-Dec-2005.) (Proof shortened by Andrew Salmon, 9-Jul-2011.)
|
⊢ (𝐴 ∈ 𝑉 → ⦋𝐴 / 𝑥⦌{𝑦 ∣ 𝜑} = {𝑦 ∣ [𝐴 / 𝑥]𝜑}) |
|
Theorem | cbvralcsf 2908 |
A more general version of cbvralf 2527 that doesn't require 𝐴 and 𝐵
to be distinct from 𝑥 or 𝑦. Changes bound
variables using
implicit substitution. (Contributed by Andrew Salmon, 13-Jul-2011.)
|
⊢ Ⅎ𝑦𝐴
& ⊢ Ⅎ𝑥𝐵
& ⊢ Ⅎ𝑦𝜑
& ⊢ Ⅎ𝑥𝜓
& ⊢ (𝑥 = 𝑦 → 𝐴 = 𝐵)
& ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝜑 ↔ ∀𝑦 ∈ 𝐵 𝜓) |
|
Theorem | cbvrexcsf 2909 |
A more general version of cbvrexf 2528 that has no distinct variable
restrictions. Changes bound variables using implicit substitution.
(Contributed by Andrew Salmon, 13-Jul-2011.) (Proof shortened by Mario
Carneiro, 7-Dec-2014.)
|
⊢ Ⅎ𝑦𝐴
& ⊢ Ⅎ𝑥𝐵
& ⊢ Ⅎ𝑦𝜑
& ⊢ Ⅎ𝑥𝜓
& ⊢ (𝑥 = 𝑦 → 𝐴 = 𝐵)
& ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃𝑥 ∈ 𝐴 𝜑 ↔ ∃𝑦 ∈ 𝐵 𝜓) |
|
Theorem | cbvreucsf 2910 |
A more general version of cbvreuv 2535 that has no distinct variable
rextrictions. Changes bound variables using implicit substitution.
(Contributed by Andrew Salmon, 13-Jul-2011.)
|
⊢ Ⅎ𝑦𝐴
& ⊢ Ⅎ𝑥𝐵
& ⊢ Ⅎ𝑦𝜑
& ⊢ Ⅎ𝑥𝜓
& ⊢ (𝑥 = 𝑦 → 𝐴 = 𝐵)
& ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃!𝑥 ∈ 𝐴 𝜑 ↔ ∃!𝑦 ∈ 𝐵 𝜓) |
|
Theorem | cbvrabcsf 2911 |
A more general version of cbvrab 2555 with no distinct variable
restrictions. (Contributed by Andrew Salmon, 13-Jul-2011.)
|
⊢ Ⅎ𝑦𝐴
& ⊢ Ⅎ𝑥𝐵
& ⊢ Ⅎ𝑦𝜑
& ⊢ Ⅎ𝑥𝜓
& ⊢ (𝑥 = 𝑦 → 𝐴 = 𝐵)
& ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ {𝑥 ∈ 𝐴 ∣ 𝜑} = {𝑦 ∈ 𝐵 ∣ 𝜓} |
|
Theorem | cbvralv2 2912* |
Rule used to change the bound variable in a restricted universal
quantifier with implicit substitution which also changes the quantifier
domain. (Contributed by David Moews, 1-May-2017.)
|
⊢ (𝑥 = 𝑦 → (𝜓 ↔ 𝜒)) & ⊢ (𝑥 = 𝑦 → 𝐴 = 𝐵) ⇒ ⊢ (∀𝑥 ∈ 𝐴 𝜓 ↔ ∀𝑦 ∈ 𝐵 𝜒) |
|
Theorem | cbvrexv2 2913* |
Rule used to change the bound variable in a restricted existential
quantifier with implicit substitution which also changes the quantifier
domain. (Contributed by David Moews, 1-May-2017.)
|
⊢ (𝑥 = 𝑦 → (𝜓 ↔ 𝜒)) & ⊢ (𝑥 = 𝑦 → 𝐴 = 𝐵) ⇒ ⊢ (∃𝑥 ∈ 𝐴 𝜓 ↔ ∃𝑦 ∈ 𝐵 𝜒) |
|
2.1.11 Define basic set operations and
relations
|
|
Syntax | cdif 2914 |
Extend class notation to include class difference (read: "𝐴 minus
𝐵").
|
class (𝐴 ∖ 𝐵) |
|
Syntax | cun 2915 |
Extend class notation to include union of two classes (read: "𝐴
union 𝐵").
|
class (𝐴 ∪ 𝐵) |
|
Syntax | cin 2916 |
Extend class notation to include the intersection of two classes (read:
"𝐴 intersect 𝐵").
|
class (𝐴 ∩ 𝐵) |
|
Syntax | wss 2917 |
Extend wff notation to include the subclass relation. This is
read "𝐴 is a subclass of 𝐵 "
or "𝐵 includes 𝐴." When
𝐴 exists as a set, it is also read
"𝐴 is a subset of 𝐵."
|
wff 𝐴 ⊆ 𝐵 |
|
Syntax | wpss 2918 |
Extend wff notation with proper subclass relation.
|
wff 𝐴 ⊊ 𝐵 |
|
Theorem | difjust 2919* |
Soundness justification theorem for df-dif 2920. (Contributed by Rodolfo
Medina, 27-Apr-2010.) (Proof shortened by Andrew Salmon,
9-Jul-2011.)
|
⊢ {𝑥 ∣ (𝑥 ∈ 𝐴 ∧ ¬ 𝑥 ∈ 𝐵)} = {𝑦 ∣ (𝑦 ∈ 𝐴 ∧ ¬ 𝑦 ∈ 𝐵)} |
|
Definition | df-dif 2920* |
Define class difference, also called relative complement. Definition
5.12 of [TakeutiZaring] p. 20.
Contrast this operation with union
(𝐴
∪ 𝐵) (df-un 2922) and intersection (𝐴 ∩ 𝐵) (df-in 2924).
Several notations are used in the literature; we chose the ∖
convention used in Definition 5.3 of [Eisenberg] p. 67 instead of the
more common minus sign to reserve the latter for later use in, e.g.,
arithmetic. We will use the terminology "𝐴 excludes 𝐵 "
to
mean 𝐴 ∖ 𝐵. We will use "𝐵 is
removed from 𝐴 " to mean
𝐴
∖ {𝐵} i.e. the
removal of an element or equivalently the
exclusion of a singleton. (Contributed by NM, 29-Apr-1994.)
|
⊢ (𝐴 ∖ 𝐵) = {𝑥 ∣ (𝑥 ∈ 𝐴 ∧ ¬ 𝑥 ∈ 𝐵)} |
|
Theorem | unjust 2921* |
Soundness justification theorem for df-un 2922. (Contributed by Rodolfo
Medina, 28-Apr-2010.) (Proof shortened by Andrew Salmon,
9-Jul-2011.)
|
⊢ {𝑥 ∣ (𝑥 ∈ 𝐴 ∨ 𝑥 ∈ 𝐵)} = {𝑦 ∣ (𝑦 ∈ 𝐴 ∨ 𝑦 ∈ 𝐵)} |
|
Definition | df-un 2922* |
Define the union of two classes. Definition 5.6 of [TakeutiZaring]
p. 16. Contrast this operation with difference (𝐴 ∖ 𝐵)
(df-dif 2920) and intersection (𝐴 ∩ 𝐵) (df-in 2924). (Contributed
by NM, 23-Aug-1993.)
|
⊢ (𝐴 ∪ 𝐵) = {𝑥 ∣ (𝑥 ∈ 𝐴 ∨ 𝑥 ∈ 𝐵)} |
|
Theorem | injust 2923* |
Soundness justification theorem for df-in 2924. (Contributed by Rodolfo
Medina, 28-Apr-2010.) (Proof shortened by Andrew Salmon,
9-Jul-2011.)
|
⊢ {𝑥 ∣ (𝑥 ∈ 𝐴 ∧ 𝑥 ∈ 𝐵)} = {𝑦 ∣ (𝑦 ∈ 𝐴 ∧ 𝑦 ∈ 𝐵)} |
|
Definition | df-in 2924* |
Define the intersection of two classes. Definition 5.6 of
[TakeutiZaring] p. 16. Contrast
this operation with union
(𝐴
∪ 𝐵) (df-un 2922) and difference (𝐴 ∖ 𝐵) (df-dif 2920).
(Contributed by NM, 29-Apr-1994.)
|
⊢ (𝐴 ∩ 𝐵) = {𝑥 ∣ (𝑥 ∈ 𝐴 ∧ 𝑥 ∈ 𝐵)} |
|
Theorem | dfin5 2925* |
Alternate definition for the intersection of two classes. (Contributed
by NM, 6-Jul-2005.)
|
⊢ (𝐴 ∩ 𝐵) = {𝑥 ∈ 𝐴 ∣ 𝑥 ∈ 𝐵} |
|
Theorem | dfdif2 2926* |
Alternate definition of class difference. (Contributed by NM,
25-Mar-2004.)
|
⊢ (𝐴 ∖ 𝐵) = {𝑥 ∈ 𝐴 ∣ ¬ 𝑥 ∈ 𝐵} |
|
Theorem | eldif 2927 |
Expansion of membership in a class difference. (Contributed by NM,
29-Apr-1994.)
|
⊢ (𝐴 ∈ (𝐵 ∖ 𝐶) ↔ (𝐴 ∈ 𝐵 ∧ ¬ 𝐴 ∈ 𝐶)) |
|
Theorem | eldifd 2928 |
If a class is in one class and not another, it is also in their
difference. One-way deduction form of eldif 2927. (Contributed by David
Moews, 1-May-2017.)
|
⊢ (𝜑 → 𝐴 ∈ 𝐵)
& ⊢ (𝜑 → ¬ 𝐴 ∈ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ∈ (𝐵 ∖ 𝐶)) |
|
Theorem | eldifad 2929 |
If a class is in the difference of two classes, it is also in the
minuend. One-way deduction form of eldif 2927. (Contributed by David
Moews, 1-May-2017.)
|
⊢ (𝜑 → 𝐴 ∈ (𝐵 ∖ 𝐶)) ⇒ ⊢ (𝜑 → 𝐴 ∈ 𝐵) |
|
Theorem | eldifbd 2930 |
If a class is in the difference of two classes, it is not in the
subtrahend. One-way deduction form of eldif 2927. (Contributed by David
Moews, 1-May-2017.)
|
⊢ (𝜑 → 𝐴 ∈ (𝐵 ∖ 𝐶)) ⇒ ⊢ (𝜑 → ¬ 𝐴 ∈ 𝐶) |
|
2.1.12 Subclasses and subsets
|
|
Definition | df-ss 2931 |
Define the subclass relationship. Exercise 9 of [TakeutiZaring] p. 18.
Note that 𝐴 ⊆ 𝐴 (proved in ssid 2964). Contrast this relationship with
the relationship 𝐴 ⊊ 𝐵 (as will be defined in df-pss 2933). For a more
traditional definition, but requiring a dummy variable, see dfss2 2934 (or
dfss3 2935 which is similar). (Contributed by NM,
27-Apr-1994.)
|
⊢ (𝐴 ⊆ 𝐵 ↔ (𝐴 ∩ 𝐵) = 𝐴) |
|
Theorem | dfss 2932 |
Variant of subclass definition df-ss 2931. (Contributed by NM,
3-Sep-2004.)
|
⊢ (𝐴 ⊆ 𝐵 ↔ 𝐴 = (𝐴 ∩ 𝐵)) |
|
Definition | df-pss 2933 |
Define proper subclass relationship between two classes. Definition 5.9
of [TakeutiZaring] p. 17. Note that
¬ 𝐴
⊊ 𝐴 (proved
in pssirr 3044).
Contrast this relationship with the relationship 𝐴 ⊆ 𝐵 (as defined in
df-ss 2931). Other possible definitions are given by dfpss2 3029 and
dfpss3 3030. (Contributed by NM, 7-Feb-1996.)
|
⊢ (𝐴 ⊊ 𝐵 ↔ (𝐴 ⊆ 𝐵 ∧ 𝐴 ≠ 𝐵)) |
|
Theorem | dfss2 2934* |
Alternate definition of the subclass relationship between two classes.
Definition 5.9 of [TakeutiZaring]
p. 17. (Contributed by NM,
8-Jan-2002.)
|
⊢ (𝐴 ⊆ 𝐵 ↔ ∀𝑥(𝑥 ∈ 𝐴 → 𝑥 ∈ 𝐵)) |
|
Theorem | dfss3 2935* |
Alternate definition of subclass relationship. (Contributed by NM,
14-Oct-1999.)
|
⊢ (𝐴 ⊆ 𝐵 ↔ ∀𝑥 ∈ 𝐴 𝑥 ∈ 𝐵) |
|
Theorem | dfss2f 2936 |
Equivalence for subclass relation, using bound-variable hypotheses
instead of distinct variable conditions. (Contributed by NM,
3-Jul-1994.) (Revised by Andrew Salmon, 27-Aug-2011.)
|
⊢ Ⅎ𝑥𝐴
& ⊢ Ⅎ𝑥𝐵 ⇒ ⊢ (𝐴 ⊆ 𝐵 ↔ ∀𝑥(𝑥 ∈ 𝐴 → 𝑥 ∈ 𝐵)) |
|
Theorem | dfss3f 2937 |
Equivalence for subclass relation, using bound-variable hypotheses
instead of distinct variable conditions. (Contributed by NM,
20-Mar-2004.)
|
⊢ Ⅎ𝑥𝐴
& ⊢ Ⅎ𝑥𝐵 ⇒ ⊢ (𝐴 ⊆ 𝐵 ↔ ∀𝑥 ∈ 𝐴 𝑥 ∈ 𝐵) |
|
Theorem | nfss 2938 |
If 𝑥 is not free in 𝐴 and 𝐵, it is
not free in 𝐴 ⊆ 𝐵.
(Contributed by NM, 27-Dec-1996.)
|
⊢ Ⅎ𝑥𝐴
& ⊢ Ⅎ𝑥𝐵 ⇒ ⊢ Ⅎ𝑥 𝐴 ⊆ 𝐵 |
|
Theorem | ssel 2939 |
Membership relationships follow from a subclass relationship.
(Contributed by NM, 5-Aug-1993.)
|
⊢ (𝐴 ⊆ 𝐵 → (𝐶 ∈ 𝐴 → 𝐶 ∈ 𝐵)) |
|
Theorem | ssel2 2940 |
Membership relationships follow from a subclass relationship.
(Contributed by NM, 7-Jun-2004.)
|
⊢ ((𝐴 ⊆ 𝐵 ∧ 𝐶 ∈ 𝐴) → 𝐶 ∈ 𝐵) |
|
Theorem | sseli 2941 |
Membership inference from subclass relationship. (Contributed by NM,
5-Aug-1993.)
|
⊢ 𝐴 ⊆ 𝐵 ⇒ ⊢ (𝐶 ∈ 𝐴 → 𝐶 ∈ 𝐵) |
|
Theorem | sselii 2942 |
Membership inference from subclass relationship. (Contributed by NM,
31-May-1999.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ 𝐶 ∈ 𝐴 ⇒ ⊢ 𝐶 ∈ 𝐵 |
|
Theorem | sseldi 2943 |
Membership inference from subclass relationship. (Contributed by NM,
25-Jun-2014.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ (𝜑 → 𝐶 ∈ 𝐴) ⇒ ⊢ (𝜑 → 𝐶 ∈ 𝐵) |
|
Theorem | sseld 2944 |
Membership deduction from subclass relationship. (Contributed by NM,
15-Nov-1995.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵) ⇒ ⊢ (𝜑 → (𝐶 ∈ 𝐴 → 𝐶 ∈ 𝐵)) |
|
Theorem | sselda 2945 |
Membership deduction from subclass relationship. (Contributed by NM,
26-Jun-2014.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵) ⇒ ⊢ ((𝜑 ∧ 𝐶 ∈ 𝐴) → 𝐶 ∈ 𝐵) |
|
Theorem | sseldd 2946 |
Membership inference from subclass relationship. (Contributed by NM,
14-Dec-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐶 ∈ 𝐴) ⇒ ⊢ (𝜑 → 𝐶 ∈ 𝐵) |
|
Theorem | ssneld 2947 |
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 2948 |
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 2949* |
Inference rule based on subclass definition. (Contributed by NM,
5-Aug-1993.)
|
⊢ (𝑥 ∈ 𝐴 → 𝑥 ∈ 𝐵) ⇒ ⊢ 𝐴 ⊆ 𝐵 |
|
Theorem | ssrd 2950 |
Deduction rule based on subclass definition. (Contributed by Thierry
Arnoux, 8-Mar-2017.)
|
⊢ Ⅎ𝑥𝜑
& ⊢ Ⅎ𝑥𝐴
& ⊢ Ⅎ𝑥𝐵
& ⊢ (𝜑 → (𝑥 ∈ 𝐴 → 𝑥 ∈ 𝐵)) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐵) |
|
Theorem | ssrdv 2951* |
Deduction rule based on subclass definition. (Contributed by NM,
15-Nov-1995.)
|
⊢ (𝜑 → (𝑥 ∈ 𝐴 → 𝑥 ∈ 𝐵)) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐵) |
|
Theorem | sstr2 2952 |
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 2953 |
Transitivity of subclasses. Theorem 6 of [Suppes] p. 23. (Contributed by
NM, 5-Sep-2003.)
|
⊢ ((𝐴 ⊆ 𝐵 ∧ 𝐵 ⊆ 𝐶) → 𝐴 ⊆ 𝐶) |
|
Theorem | sstri 2954 |
Subclass transitivity inference. (Contributed by NM, 5-May-2000.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ 𝐵 ⊆ 𝐶 ⇒ ⊢ 𝐴 ⊆ 𝐶 |
|
Theorem | sstrd 2955 |
Subclass transitivity deduction. (Contributed by NM, 2-Jun-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl5ss 2956 |
Subclass transitivity deduction. (Contributed by NM, 6-Feb-2014.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl6ss 2957 |
Subclass transitivity deduction. (Contributed by Jonathan Ben-Naim,
3-Jun-2011.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ 𝐵 ⊆ 𝐶 ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | sylan9ss 2958 |
A subclass transitivity deduction. (Contributed by NM, 27-Sep-2004.)
(Proof shortened by Andrew Salmon, 14-Jun-2011.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜓 → 𝐵 ⊆ 𝐶) ⇒ ⊢ ((𝜑 ∧ 𝜓) → 𝐴 ⊆ 𝐶) |
|
Theorem | sylan9ssr 2959 |
A subclass transitivity deduction. (Contributed by NM, 27-Sep-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜓 → 𝐵 ⊆ 𝐶) ⇒ ⊢ ((𝜓 ∧ 𝜑) → 𝐴 ⊆ 𝐶) |
|
Theorem | eqss 2960 |
The subclass relationship is antisymmetric. Compare Theorem 4 of
[Suppes] p. 22. (Contributed by NM,
5-Aug-1993.)
|
⊢ (𝐴 = 𝐵 ↔ (𝐴 ⊆ 𝐵 ∧ 𝐵 ⊆ 𝐴)) |
|
Theorem | eqssi 2961 |
Infer equality from two subclass relationships. Compare Theorem 4 of
[Suppes] p. 22. (Contributed by NM,
9-Sep-1993.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ 𝐵 ⊆ 𝐴 ⇒ ⊢ 𝐴 = 𝐵 |
|
Theorem | eqssd 2962 |
Equality deduction from two subclass relationships. Compare Theorem 4
of [Suppes] p. 22. (Contributed by NM,
27-Jun-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐵 ⊆ 𝐴) ⇒ ⊢ (𝜑 → 𝐴 = 𝐵) |
|
Theorem | eqrd 2963 |
Deduce equality of classes from equivalence of membership. (Contributed
by Thierry Arnoux, 21-Mar-2017.)
|
⊢ Ⅎ𝑥𝜑
& ⊢ Ⅎ𝑥𝐴
& ⊢ Ⅎ𝑥𝐵
& ⊢ (𝜑 → (𝑥 ∈ 𝐴 ↔ 𝑥 ∈ 𝐵)) ⇒ ⊢ (𝜑 → 𝐴 = 𝐵) |
|
Theorem | ssid 2964 |
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 2965 |
Any class is a subclass of the universal class. (Contributed by NM,
31-Oct-1995.)
|
⊢ 𝐴 ⊆ V |
|
Theorem | sseq1 2966 |
Equality theorem for subclasses. (Contributed by NM, 5-Aug-1993.) (Proof
shortened by Andrew Salmon, 21-Jun-2011.)
|
⊢ (𝐴 = 𝐵 → (𝐴 ⊆ 𝐶 ↔ 𝐵 ⊆ 𝐶)) |
|
Theorem | sseq2 2967 |
Equality theorem for the subclass relationship. (Contributed by NM,
25-Jun-1998.)
|
⊢ (𝐴 = 𝐵 → (𝐶 ⊆ 𝐴 ↔ 𝐶 ⊆ 𝐵)) |
|
Theorem | sseq12 2968 |
Equality theorem for the subclass relationship. (Contributed by NM,
31-May-1999.)
|
⊢ ((𝐴 = 𝐵 ∧ 𝐶 = 𝐷) → (𝐴 ⊆ 𝐶 ↔ 𝐵 ⊆ 𝐷)) |
|
Theorem | sseq1i 2969 |
An equality inference for the subclass relationship. (Contributed by
NM, 18-Aug-1993.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ (𝐴 ⊆ 𝐶 ↔ 𝐵 ⊆ 𝐶) |
|
Theorem | sseq2i 2970 |
An equality inference for the subclass relationship. (Contributed by
NM, 30-Aug-1993.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ (𝐶 ⊆ 𝐴 ↔ 𝐶 ⊆ 𝐵) |
|
Theorem | sseq12i 2971 |
An equality inference for the subclass relationship. (Contributed by
NM, 31-May-1999.) (Proof shortened by Eric Schmidt, 26-Jan-2007.)
|
⊢ 𝐴 = 𝐵
& ⊢ 𝐶 = 𝐷 ⇒ ⊢ (𝐴 ⊆ 𝐶 ↔ 𝐵 ⊆ 𝐷) |
|
Theorem | sseq1d 2972 |
An equality deduction for the subclass relationship. (Contributed by
NM, 14-Aug-1994.)
|
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐴 ⊆ 𝐶 ↔ 𝐵 ⊆ 𝐶)) |
|
Theorem | sseq2d 2973 |
An equality deduction for the subclass relationship. (Contributed by
NM, 14-Aug-1994.)
|
⊢ (𝜑 → 𝐴 = 𝐵) ⇒ ⊢ (𝜑 → (𝐶 ⊆ 𝐴 ↔ 𝐶 ⊆ 𝐵)) |
|
Theorem | sseq12d 2974 |
An equality deduction for the subclass relationship. (Contributed by
NM, 31-May-1999.)
|
⊢ (𝜑 → 𝐴 = 𝐵)
& ⊢ (𝜑 → 𝐶 = 𝐷) ⇒ ⊢ (𝜑 → (𝐴 ⊆ 𝐶 ↔ 𝐵 ⊆ 𝐷)) |
|
Theorem | eqsstri 2975 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 16-Jul-1995.)
|
⊢ 𝐴 = 𝐵
& ⊢ 𝐵 ⊆ 𝐶 ⇒ ⊢ 𝐴 ⊆ 𝐶 |
|
Theorem | eqsstr3i 2976 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 19-Oct-1999.)
|
⊢ 𝐵 = 𝐴
& ⊢ 𝐵 ⊆ 𝐶 ⇒ ⊢ 𝐴 ⊆ 𝐶 |
|
Theorem | sseqtri 2977 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 28-Jul-1995.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ 𝐵 = 𝐶 ⇒ ⊢ 𝐴 ⊆ 𝐶 |
|
Theorem | sseqtr4i 2978 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 4-Apr-1995.)
|
⊢ 𝐴 ⊆ 𝐵
& ⊢ 𝐶 = 𝐵 ⇒ ⊢ 𝐴 ⊆ 𝐶 |
|
Theorem | eqsstrd 2979 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 25-Apr-2004.)
|
⊢ (𝜑 → 𝐴 = 𝐵)
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | eqsstr3d 2980 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 25-Apr-2004.)
|
⊢ (𝜑 → 𝐵 = 𝐴)
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | sseqtrd 2981 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 25-Apr-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐵 = 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | sseqtr4d 2982 |
Substitution of equality into a subclass relationship. (Contributed by
NM, 25-Apr-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐶 = 𝐵) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | 3sstr3i 2983 |
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 2984 |
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 2985 |
Substitution of equality into both sides of a subclass relationship.
(Contributed by NM, 1-Oct-2000.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ 𝐴 = 𝐶
& ⊢ 𝐵 = 𝐷 ⇒ ⊢ (𝜑 → 𝐶 ⊆ 𝐷) |
|
Theorem | 3sstr4g 2986 |
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 2987 |
Substitution of equality into both sides of a subclass relationship.
(Contributed by NM, 1-Oct-2000.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ (𝜑 → 𝐴 = 𝐶)
& ⊢ (𝜑 → 𝐵 = 𝐷) ⇒ ⊢ (𝜑 → 𝐶 ⊆ 𝐷) |
|
Theorem | 3sstr4d 2988 |
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 2989 |
B chained subclass and equality deduction. (Contributed by NM,
25-Apr-2004.)
|
⊢ 𝐴 = 𝐵
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl5eqssr 2990 |
B chained subclass and equality deduction. (Contributed by NM,
25-Apr-2004.)
|
⊢ 𝐵 = 𝐴
& ⊢ (𝜑 → 𝐵 ⊆ 𝐶) ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl6sseq 2991 |
A chained subclass and equality deduction. (Contributed by NM,
25-Apr-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ 𝐵 = 𝐶 ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl6sseqr 2992 |
A chained subclass and equality deduction. (Contributed by NM,
25-Apr-2004.)
|
⊢ (𝜑 → 𝐴 ⊆ 𝐵)
& ⊢ 𝐶 = 𝐵 ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl5sseq 2993 |
Subclass transitivity deduction. (Contributed by Jonathan Ben-Naim,
3-Jun-2011.)
|
⊢ 𝐵 ⊆ 𝐴
& ⊢ (𝜑 → 𝐴 = 𝐶) ⇒ ⊢ (𝜑 → 𝐵 ⊆ 𝐶) |
|
Theorem | syl5sseqr 2994 |
Subclass transitivity deduction. (Contributed by Jonathan Ben-Naim,
3-Jun-2011.)
|
⊢ 𝐵 ⊆ 𝐴
& ⊢ (𝜑 → 𝐶 = 𝐴) ⇒ ⊢ (𝜑 → 𝐵 ⊆ 𝐶) |
|
Theorem | syl6eqss 2995 |
A chained subclass and equality deduction. (Contributed by Mario
Carneiro, 2-Jan-2017.)
|
⊢ (𝜑 → 𝐴 = 𝐵)
& ⊢ 𝐵 ⊆ 𝐶 ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | syl6eqssr 2996 |
A chained subclass and equality deduction. (Contributed by Mario
Carneiro, 2-Jan-2017.)
|
⊢ (𝜑 → 𝐵 = 𝐴)
& ⊢ 𝐵 ⊆ 𝐶 ⇒ ⊢ (𝜑 → 𝐴 ⊆ 𝐶) |
|
Theorem | eqimss 2997 |
Equality implies the subclass relation. (Contributed by NM, 5-Aug-1993.)
(Proof shortened by Andrew Salmon, 21-Jun-2011.)
|
⊢ (𝐴 = 𝐵 → 𝐴 ⊆ 𝐵) |
|
Theorem | eqimss2 2998 |
Equality implies the subclass relation. (Contributed by NM,
23-Nov-2003.)
|
⊢ (𝐵 = 𝐴 → 𝐴 ⊆ 𝐵) |
|
Theorem | eqimssi 2999 |
Infer subclass relationship from equality. (Contributed by NM,
6-Jan-2007.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ 𝐴 ⊆ 𝐵 |
|
Theorem | eqimss2i 3000 |
Infer subclass relationship from equality. (Contributed by NM,
7-Jan-2007.)
|
⊢ 𝐴 = 𝐵 ⇒ ⊢ 𝐵 ⊆ 𝐴 |