Theorem List for Metamath Proof Explorer - 39001-39100 *Has distinct variable
group(s)
Type | Label | Description |
Statement |
|
Theorem | exbir 39001 |
Exportation implication also converting the consequent from a
biconditional to an implication. Derived automatically from exbirVD 39402.
(Contributed by Alan Sare, 31-Dec-2011.)
|
⊢ (((𝜑 ∧ 𝜓) → (𝜒 ↔ 𝜃)) → (𝜑 → (𝜓 → (𝜃 → 𝜒)))) |
|
Theorem | 3impexpbicom 39002 |
Version of 3impexp 1311 where in addition the consequent is commuted.
(Contributed by Alan Sare, 31-Dec-2011.)
|
⊢ (((𝜑 ∧ 𝜓 ∧ 𝜒) → (𝜃 ↔ 𝜏)) ↔ (𝜑 → (𝜓 → (𝜒 → (𝜏 ↔ 𝜃))))) |
|
Theorem | 3impexpbicomi 39003 |
Inference associated with 3impexpbicom 39002. Derived automatically from
3impexpbicomiVD 39407. (Contributed by Alan Sare, 31-Dec-2011.)
|
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → (𝜃 ↔ 𝜏)) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 → (𝜏 ↔ 𝜃)))) |
|
20.31.2 Supplementary unification
deductions
|
|
Theorem | bi1imp 39004 |
Importation inference similar to imp 444, except the outermost
implication of the hypothesis is a biconditional. (Contributed by Alan
Sare, 6-Nov-2017.)
|
⊢ (𝜑 ↔ (𝜓 → 𝜒)) ⇒ ⊢ ((𝜑 ∧ 𝜓) → 𝜒) |
|
Theorem | bi2imp 39005 |
Importation inference similar to imp 444, except both implications of the
hypothesis are biconditionals. (Contributed by Alan Sare,
6-Nov-2017.)
|
⊢ (𝜑 ↔ (𝜓 ↔ 𝜒)) ⇒ ⊢ ((𝜑 ∧ 𝜓) → 𝜒) |
|
Theorem | bi3impb 39006 |
Similar to 3impb 1279 with implication in hypothesis replaced by
biconditional. (Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ ((𝜑 ∧ (𝜓 ∧ 𝜒)) ↔ 𝜃) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi3impa 39007 |
Similar to 3impa 1278 with implication in hypothesis replaced by
biconditional. (Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ (((𝜑 ∧ 𝜓) ∧ 𝜒) ↔ 𝜃) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi23impib 39008 |
3impib 1281 with the inner implication of the hypothesis
a biconditional.
(Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ (𝜑 → ((𝜓 ∧ 𝜒) ↔ 𝜃)) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi13impib 39009 |
3impib 1281 with the outer implication of the hypothesis
a biconditional.
(Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ (𝜑 ↔ ((𝜓 ∧ 𝜒) → 𝜃)) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi123impib 39010 |
3impib 1281 with the implications of the hypothesis
biconditionals.
(Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ (𝜑 ↔ ((𝜓 ∧ 𝜒) ↔ 𝜃)) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi13impia 39011 |
3impia 1280 with the outer implication of the hypothesis
a biconditional.
(Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ ((𝜑 ∧ 𝜓) ↔ (𝜒 → 𝜃)) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi123impia 39012 |
3impia 1280 with the implications of the hypothesis
biconditionals.
(Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ ((𝜑 ∧ 𝜓) ↔ (𝜒 ↔ 𝜃)) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi33imp12 39013 |
3imp 1275 with innermost implication of the hypothesis
a biconditional.
(Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ (𝜑 → (𝜓 → (𝜒 ↔ 𝜃))) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi23imp13 39014 |
3imp 1275 with middle implication of the hypothesis a
biconditional.
(Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ (𝜑 → (𝜓 ↔ (𝜒 → 𝜃))) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi13imp23 39015 |
3imp 1275 with outermost implication of the hypothesis
a biconditional.
(Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ (𝜑 ↔ (𝜓 → (𝜒 → 𝜃))) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi13imp2 39016 |
Similar to 3imp 1275 except the outermost and innermost
implications are
biconditionals. (Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ (𝜑 ↔ (𝜓 → (𝜒 ↔ 𝜃))) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi12imp3 39017 |
Similar to 3imp 1275 except all but innermost implication are
biconditionals. (Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ (𝜑 ↔ (𝜓 ↔ (𝜒 → 𝜃))) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi23imp1 39018 |
Similar to 3imp 1275 except all but outermost implication are
biconditionals. (Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ (𝜑 → (𝜓 ↔ (𝜒 ↔ 𝜃))) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | bi123imp0 39019 |
Similar to 3imp 1275 except all implications are biconditionals.
(Contributed by Alan Sare, 6-Nov-2017.)
|
⊢ (𝜑 ↔ (𝜓 ↔ (𝜒 ↔ 𝜃))) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝜃) |
|
Theorem | 4animp1 39020 |
A single hypothesis unification deduction with an assertion which is an
implication with a 4-right-nested conjunction antecedent. (Contributed
by Alan Sare, 30-May-2018.)
|
⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → (𝜏 ↔ 𝜃)) ⇒ ⊢ ((((𝜑 ∧ 𝜓) ∧ 𝜒) ∧ 𝜃) → 𝜏) |
|
Theorem | 4an31 39021 |
A rearrangement of conjuncts for a 4-right-nested conjunction.
(Contributed by Alan Sare, 30-May-2018.)
|
⊢ ((((𝜒 ∧ 𝜓) ∧ 𝜑) ∧ 𝜃) → 𝜏) ⇒ ⊢ ((((𝜑 ∧ 𝜓) ∧ 𝜒) ∧ 𝜃) → 𝜏) |
|
Theorem | 4an4132 39022 |
A rearrangement of conjuncts for a 4-right-nested conjunction.
(Contributed by Alan Sare, 30-May-2018.)
|
⊢ ((((𝜃 ∧ 𝜒) ∧ 𝜓) ∧ 𝜑) → 𝜏) ⇒ ⊢ ((((𝜑 ∧ 𝜓) ∧ 𝜒) ∧ 𝜃) → 𝜏) |
|
Theorem | expcomdg 39023 |
Biconditional form of expcomd 453. (Contributed by Alan Sare,
22-Jul-2012.) (New usage is discouraged.)
|
⊢ ((𝜑 → ((𝜓 ∧ 𝜒) → 𝜃)) ↔ (𝜑 → (𝜒 → (𝜓 → 𝜃)))) |
|
20.31.3 Conventional Metamath proofs, some
derived from VD proofs
|
|
Theorem | iidn3 39024 |
idn3 39157 without virtual deduction connectives.
Special theorem needed for
the Virtual Deduction translation tool. (Contributed by Alan Sare,
23-Jul-2011.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜑 → (𝜓 → (𝜒 → 𝜒))) |
|
Theorem | ee222 39025 |
e222 39178 without virtual deduction connectives.
Special theorem needed
for the Virtual Deduction translation tool. (Contributed by Alan Sare,
7-Jul-2011.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → (𝜓 → 𝜃)) & ⊢ (𝜑 → (𝜓 → 𝜏)) & ⊢ (𝜒 → (𝜃 → (𝜏 → 𝜂))) ⇒ ⊢ (𝜑 → (𝜓 → 𝜂)) |
|
Theorem | ee3bir 39026 |
Right-biconditional form of e3 39281 without virtual deduction connectives.
Special theorem needed for the Virtual Deduction translation tool.
(Contributed by Alan Sare, 22-Jul-2011.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (𝜑 → (𝜓 → (𝜒 → 𝜃))) & ⊢ (𝜏 ↔ 𝜃) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 → 𝜏))) |
|
Theorem | ee13 39027 |
e13 39292 without virtual deduction connectives.
Special theorem needed for
the Virtual Deduction translation tool. (Contributed by Alan Sare,
28-Oct-2011.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜑 → 𝜓)
& ⊢ (𝜑 → (𝜒 → (𝜃 → 𝜏))) & ⊢ (𝜓 → (𝜏 → 𝜂)) ⇒ ⊢ (𝜑 → (𝜒 → (𝜃 → 𝜂))) |
|
Theorem | ee121 39028 |
e121 39198 without virtual deductions. (Contributed by
Alan Sare,
13-Jul-2011.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜑 → 𝜓)
& ⊢ (𝜑 → (𝜒 → 𝜃)) & ⊢ (𝜑 → 𝜏)
& ⊢ (𝜓 → (𝜃 → (𝜏 → 𝜂))) ⇒ ⊢ (𝜑 → (𝜒 → 𝜂)) |
|
Theorem | ee122 39029 |
e122 39195 without virtual deductions. (Contributed by
Alan Sare,
13-Jul-2011.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜑 → 𝜓)
& ⊢ (𝜑 → (𝜒 → 𝜃)) & ⊢ (𝜑 → (𝜒 → 𝜏)) & ⊢ (𝜓 → (𝜃 → (𝜏 → 𝜂))) ⇒ ⊢ (𝜑 → (𝜒 → 𝜂)) |
|
Theorem | ee333 39030 |
e333 39277 without virtual deductions. (Contributed by
Alan Sare,
17-Jul-2011.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜑 → (𝜓 → (𝜒 → 𝜃))) & ⊢ (𝜑 → (𝜓 → (𝜒 → 𝜏))) & ⊢ (𝜑 → (𝜓 → (𝜒 → 𝜂))) & ⊢ (𝜃 → (𝜏 → (𝜂 → 𝜁))) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 → 𝜁))) |
|
Theorem | ee323 39031 |
e323 39310 without virtual deductions. (Contributed by
Alan Sare,
17-Apr-2012.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜑 → (𝜓 → (𝜒 → 𝜃))) & ⊢ (𝜑 → (𝜓 → 𝜏)) & ⊢ (𝜑 → (𝜓 → (𝜒 → 𝜂))) & ⊢ (𝜃 → (𝜏 → (𝜂 → 𝜁))) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 → 𝜁))) |
|
Theorem | 3ornot23 39032 |
If the second and third disjuncts of a true triple disjunction are false,
then the first disjunct is true. Automatically derived from
3ornot23VD 39396. (Contributed by Alan Sare, 31-Dec-2011.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ ((¬ 𝜑 ∧ ¬ 𝜓) → ((𝜒 ∨ 𝜑 ∨ 𝜓) → 𝜒)) |
|
Theorem | orbi1r 39033 |
orbi1 742 with order of disjuncts reversed. Derived
from orbi1rVD 39397.
(Contributed by Alan Sare, 31-Dec-2011.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ ((𝜑 ↔ 𝜓) → ((𝜒 ∨ 𝜑) ↔ (𝜒 ∨ 𝜓))) |
|
Theorem | 3orbi123 39034 |
pm4.39 933 with a 3-conjunct antecedent. This proof is
3orbi123VD 39399
automatically translated and minimized. (Contributed by Alan Sare,
31-Dec-2011.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (((𝜑 ↔ 𝜓) ∧ (𝜒 ↔ 𝜃) ∧ (𝜏 ↔ 𝜂)) → ((𝜑 ∨ 𝜒 ∨ 𝜏) ↔ (𝜓 ∨ 𝜃 ∨ 𝜂))) |
|
Theorem | syl5imp 39035 |
Closed form of syl5 34. Derived automatically from syl5impVD 39413.
(Contributed by Alan Sare, 31-Dec-2011.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ ((𝜑 → (𝜓 → 𝜒)) → ((𝜃 → 𝜓) → (𝜑 → (𝜃 → 𝜒)))) |
|
Theorem | impexpd 39036 |
The following User's Proof is a Virtual Deduction proof completed
automatically by the tools program completeusersproof.cmd, which invokes
Mel L. O'Cat's mmj2 and Norm Megill's Metamath Proof Assistant. After the
User's Proof was completed, it was minimized. The completed User's Proof
before minimization is not shown. (Contributed by Alan Sare,
18-Mar-2012.) (Proof modification is discouraged.)
(New usage is discouraged.)
1:: | ⊢ (((𝜓 ∧ 𝜒) → 𝜃) ↔ (𝜓 → (𝜒 →
𝜃)))
| qed:1: | ⊢ ((𝜑 → ((𝜓 ∧ 𝜒) → 𝜃)) ↔ (𝜑
→ (𝜓 → (𝜒 → 𝜃))))
|
|
⊢ ((𝜑 → ((𝜓 ∧ 𝜒) → 𝜃)) ↔ (𝜑 → (𝜓 → (𝜒 → 𝜃)))) |
|
Theorem | com3rgbi 39037 |
The following User's Proof is a Virtual Deduction proof completed
automatically by the tools program completeusersproof.cmd, which invokes
Mel L. O'Cat's mmj2 and Norm Megill's Metamath Proof Assistant. The
completed Virtual Deduction Proof (not shown) was minimized. The
minimized proof is shown.
(Contributed by Alan Sare, 18-Mar-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
1:: | ⊢ ((𝜑 → (𝜓 → (𝜒 → 𝜃)))
→ (𝜑 → (𝜒 → (𝜓 → 𝜃))))
| 2:: | ⊢ ((𝜑 → (𝜒 → (𝜓 → 𝜃)))
→ (𝜒 → (𝜑 → (𝜓 → 𝜃))))
| 3:1,2: | ⊢ ((𝜑 → (𝜓 → (𝜒 → 𝜃)))
→ (𝜒 → (𝜑 → (𝜓 → 𝜃))))
| 4:: | ⊢ ((𝜒 → (𝜑 → (𝜓 → 𝜃)))
→ (𝜑 → (𝜒 → (𝜓 → 𝜃))))
| 5:: | ⊢ ((𝜑 → (𝜒 → (𝜓 → 𝜃)))
→ (𝜑 → (𝜓 → (𝜒 → 𝜃))))
| 6:4,5: | ⊢ ((𝜒 → (𝜑 → (𝜓 → 𝜃)))
→ (𝜑 → (𝜓 → (𝜒 → 𝜃))))
| qed:3,6: | ⊢ ((𝜑 → (𝜓 → (𝜒 → 𝜃)))
↔ (𝜒 → (𝜑 → (𝜓 → 𝜃))))
|
|
⊢ ((𝜑 → (𝜓 → (𝜒 → 𝜃))) ↔ (𝜒 → (𝜑 → (𝜓 → 𝜃)))) |
|
Theorem | impexpdcom 39038 |
The following User's Proof is a Virtual Deduction proof completed
automatically by the tools program completeusersproof.cmd, which invokes
Mel L. O'Cat's mmj2 and Norm Megill's Metamath Proof Assistant. The
completed Virtual Deduction Proof (not shown) was minimized. The
minimized proof is shown. (Contributed by Alan Sare, 18-Mar-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
1:: | ⊢ ((𝜑 → ((𝜓 ∧ 𝜒) → 𝜃))
↔ (𝜑 → (𝜓 → (𝜒 → 𝜃))))
| 2:: | ⊢ ((𝜓 → (𝜒 → (𝜑 → 𝜃)))
↔ (𝜑 → (𝜓 → (𝜒 → 𝜃))))
| qed:1,2: | ⊢ ((𝜑 → ((𝜓 ∧ 𝜒) → 𝜃))
↔ (𝜓 → (𝜒 → (𝜑 → 𝜃))))
|
|
⊢ ((𝜑 → ((𝜓 ∧ 𝜒) → 𝜃)) ↔ (𝜓 → (𝜒 → (𝜑 → 𝜃)))) |
|
Theorem | ee1111 39039 |
Non-virtual deduction form of e1111 39217. (Contributed by Alan Sare,
18-Mar-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
The following User's Proof is a Virtual Deduction proof
completed automatically by the tools program completeusersproof.cmd,
which invokes Mel L. O'Cat's mmj2 and Norm Megill's Metamath Proof
Assistant. The completed Virtual Deduction Proof (not shown) was
minimized. The minimized proof is shown.
h1:: | ⊢ (𝜑 → 𝜓)
| h2:: | ⊢ (𝜑 → 𝜒)
| h3:: | ⊢ (𝜑 → 𝜃)
| h4:: | ⊢ (𝜑 → 𝜏)
| h5:: | ⊢ (𝜓 → (𝜒 → (𝜃 → (𝜏 → 𝜂))))
| 6:1,5: | ⊢ (𝜑 → (𝜒 → (𝜃 → (𝜏 → 𝜂))))
| 7:6: | ⊢ (𝜒 → (𝜑 → (𝜃 → (𝜏 → 𝜂))))
| 8:2,7: | ⊢ (𝜑 → (𝜑 → (𝜃 → (𝜏 → 𝜂))))
| 9:8: | ⊢ (𝜑 → (𝜃 → (𝜏 → 𝜂)))
| 10:9: | ⊢ (𝜃 → (𝜑 → (𝜏 → 𝜂)))
| 11:3,10: | ⊢ (𝜑 → (𝜑 → (𝜏 → 𝜂)))
| 12:11: | ⊢ (𝜑 → (𝜏 → 𝜂))
| 13:12: | ⊢ (𝜏 → (𝜑 → 𝜂))
| 14:4,13: | ⊢ (𝜑 → (𝜑 → 𝜂))
| qed:14: | ⊢ (𝜑 → 𝜂)
|
|
⊢ (𝜑 → 𝜓)
& ⊢ (𝜑 → 𝜒)
& ⊢ (𝜑 → 𝜃)
& ⊢ (𝜑 → 𝜏)
& ⊢ (𝜓 → (𝜒 → (𝜃 → (𝜏 → 𝜂)))) ⇒ ⊢ (𝜑 → 𝜂) |
|
Theorem | pm2.43bgbi 39040 |
Logical equivalence of a 2-left-nested implication and a 1-left-nested
implicated
when two antecedents of the former implication are identical.
(Contributed by Alan Sare, 18-Mar-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
The following User's Proof is a Virtual Deduction proof completed
automatically by the tools program completeusersproof.cmd, which invokes
Mel L. O'Cat's mmj2 and Norm Megill's Metamath Proof Assistant. The
completed Virtual
Deduction Proof (not shown) was minimized. The minimized proof is
shown.
1:: | ⊢ ((𝜑 → (𝜓 → (𝜑 → 𝜒)))
→ (𝜑 → (𝜑 → (𝜓 → 𝜒))))
| 2:: | ⊢ ((𝜑 → (𝜑 → (𝜓 → 𝜒)))
→ (𝜑 → (𝜓 → 𝜒)))
| 3:1,2: | ⊢ ((𝜑 → (𝜓 → (𝜑 → 𝜒)))
→ (𝜑 → (𝜓 → 𝜒)))
| 4:: | ⊢ ((𝜑 → (𝜓 → 𝜒)) → (𝜓
→ (𝜑 → 𝜒)))
| 5:3,4: | ⊢ ((𝜑 → (𝜓 → (𝜑 → 𝜒)))
→ (𝜓 → (𝜑 → 𝜒)))
| 6:: | ⊢ ((𝜓 → (𝜑 → 𝜒)) → (𝜑
→ (𝜓 → (𝜑 → 𝜒))))
| qed:5,6: | ⊢ ((𝜑 → (𝜓 → (𝜑 → 𝜒)))
↔ (𝜓 → (𝜑 → 𝜒)))
|
|
⊢ ((𝜑 → (𝜓 → (𝜑 → 𝜒))) ↔ (𝜓 → (𝜑 → 𝜒))) |
|
Theorem | pm2.43cbi 39041 |
Logical equivalence of a 3-left-nested implication and a 2-left-nested
implicated when two antecedents of the former implication are identical.
(Contributed by Alan Sare, 18-Mar-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
The following User's Proof is
a Virtual Deduction proof completed automatically by the tools program
completeusersproof.cmd, which invokes Mel L. O'Cat's mmj2 and Norm
Megill's Metamath Proof Assistant. The completed Virtual Deduction Proof
(not shown) was minimized. The minimized proof is shown.
1:: | ⊢ ((𝜑 → (𝜓 → (𝜒 → (𝜑 → 𝜃)))
) → (𝜑 → (𝜓 → (𝜑 → (𝜒 → 𝜃)))))
| 2:: | ⊢ ((𝜑 → (𝜓 → (𝜑 → (𝜒 → 𝜃)))
) → (𝜓 → (𝜑 → (𝜒 → 𝜃))))
| 3:1,2: | ⊢ ((𝜑 → (𝜓 → (𝜒 → (𝜑 → 𝜃)))
) → (𝜓 → (𝜑 → (𝜒 → 𝜃))))
| 4:: | ⊢ ((𝜓 → (𝜑 → (𝜒 → 𝜃)))
→ (𝜓 → (𝜒 → (𝜑 → 𝜃))))
| 5:3,4: | ⊢ ((𝜑 → (𝜓 → (𝜒 → (𝜑 → 𝜃)))
) → (𝜓 → (𝜒 → (𝜑 → 𝜃))))
| 6:: | ⊢ ((𝜓 → (𝜒 → (𝜑 → 𝜃)))
→ (𝜑 → (𝜓 → (𝜒 → (𝜑 → 𝜃)))))
| qed:5,6: | ⊢ ((𝜑 → (𝜓 → (𝜒 → (𝜑 → 𝜃)))
) ↔ (𝜓 → (𝜒 → (𝜑 → 𝜃))))
|
|
⊢ ((𝜑 → (𝜓 → (𝜒 → (𝜑 → 𝜃)))) ↔ (𝜓 → (𝜒 → (𝜑 → 𝜃)))) |
|
Theorem | ee233 39042 |
Non-virtual deduction form of e233 39309. (Contributed by Alan Sare,
18-Mar-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
The following User's Proof is a Virtual Deduction proof completed
automatically by the tools program completeusersproof.cmd, which invokes
Mel L. O'Cat's mmj2 and Norm Megill's Metamath Proof Assistant. The
completed Virtual
Deduction Proof (not shown) was minimized. The minimized proof is
shown.
h1:: | ⊢ (𝜑 → (𝜓 → 𝜒))
| h2:: | ⊢ (𝜑 → (𝜓 → (𝜃 → 𝜏)))
| h3:: | ⊢ (𝜑 → (𝜓 → (𝜃 → 𝜂)))
| h4:: | ⊢ (𝜒 → (𝜏 → (𝜂 → 𝜁)))
| 5:1,4: | ⊢ (𝜑 → (𝜓 → (𝜏 → (𝜂 → 𝜁)))
)
| 6:5: | ⊢ (𝜏 → (𝜑 → (𝜓 → (𝜂 → 𝜁)))
)
| 7:2,6: | ⊢ (𝜑 → (𝜓 → (𝜃 → (𝜑 → (𝜓
→ (𝜂 → 𝜁))))))
| 8:7: | ⊢ (𝜓 → (𝜃 → (𝜑 → (𝜓 → (𝜂
→ 𝜁)))))
| 9:8: | ⊢ (𝜃 → (𝜑 → (𝜓 → (𝜂 → 𝜁)))
)
| 10:9: | ⊢ (𝜑 → (𝜓 → (𝜃 → (𝜂 → 𝜁)))
)
| 11:10: | ⊢ (𝜂 → (𝜑 → (𝜓 → (𝜃 → 𝜁)))
)
| 12:3,11: | ⊢ (𝜑 → (𝜓 → (𝜃 → (𝜑 → (𝜓
→ (𝜃 → 𝜁))))))
| 13:12: | ⊢ (𝜓 → (𝜃 → (𝜑 → (𝜓 → (𝜃
→ 𝜁)))))
| 14:13: | ⊢ (𝜃 → (𝜑 → (𝜓 → (𝜃 → 𝜁)))
)
| qed:14: | ⊢ (𝜑 → (𝜓 → (𝜃 → 𝜁)))
|
|
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → (𝜓 → (𝜃 → 𝜏))) & ⊢ (𝜑 → (𝜓 → (𝜃 → 𝜂))) & ⊢ (𝜒 → (𝜏 → (𝜂 → 𝜁))) ⇒ ⊢ (𝜑 → (𝜓 → (𝜃 → 𝜁))) |
|
Theorem | imbi13 39043 |
Join three logical equivalences to form equivalence of implications.
imbi13 39043 is imbi13VD 39424 without virtual deductions and was
automatically
derived from imbi13VD 39424 using the tools program
translate..without..overwriting.cmd and Metamath's minimize command.
(Contributed by Alan Sare, 18-Mar-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ ((𝜑 ↔ 𝜓) → ((𝜒 ↔ 𝜃) → ((𝜏 ↔ 𝜂) → ((𝜑 → (𝜒 → 𝜏)) ↔ (𝜓 → (𝜃 → 𝜂)))))) |
|
Theorem | ee33 39044 |
Non-virtual deduction form of e33 39278. (Contributed by Alan Sare,
18-Mar-2012.) (Proof modification is discouraged.)
(New usage is discouraged.)
The following User's Proof is a Virtual Deduction proof
completed automatically by the tools program completeusersproof.cmd,
which invokes Mel L. O'Cat's mmj2 and Norm Megill's Metamath Proof
Assistant. The completed Virtual Deduction Proof (not shown) was
minimized. The minimized proof is shown.
h1:: | ⊢ (𝜑 → (𝜓 → (𝜒 → 𝜃)))
| h2:: | ⊢ (𝜑 → (𝜓 → (𝜒 → 𝜏)))
| h3:: | ⊢ (𝜃 → (𝜏 → 𝜂))
| 4:1,3: | ⊢ (𝜑 → (𝜓 → (𝜒 → (𝜏 → 𝜂))))
| 5:4: | ⊢ (𝜏 → (𝜑 → (𝜓 → (𝜒 → 𝜂))))
| 6:2,5: | ⊢ (𝜑 → (𝜓 → (𝜒 → (𝜑 → (𝜓 →
(𝜒 → 𝜂))))))
| 7:6: | ⊢ (𝜓 → (𝜒 → (𝜑 → (𝜓 → (𝜒 →
𝜂)))))
| 8:7: | ⊢ (𝜒 → (𝜑 → (𝜓 → (𝜒 → 𝜂))))
| qed:8: | ⊢ (𝜑 → (𝜓 → (𝜒 → 𝜂)))
|
|
⊢ (𝜑 → (𝜓 → (𝜒 → 𝜃))) & ⊢ (𝜑 → (𝜓 → (𝜒 → 𝜏))) & ⊢ (𝜃 → (𝜏 → 𝜂)) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 → 𝜂))) |
|
Theorem | con5 39045 |
Biconditional contraposition variation. This proof is con5VD 39450
automatically translated and minimized. (Contributed by Alan Sare,
21-Apr-2013.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ ((𝜑 ↔ ¬ 𝜓) → (¬ 𝜑 → 𝜓)) |
|
Theorem | con5i 39046 |
Inference form of con5 39045. (Contributed by Alan Sare, 21-Apr-2013.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (𝜑 ↔ ¬ 𝜓) ⇒ ⊢ (¬ 𝜑 → 𝜓) |
|
Theorem | exlimexi 39047 |
Inference similar to Theorem 19.23 of [Margaris] p. 90. (Contributed by
Alan Sare, 21-Apr-2013.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜓 → ∀𝑥𝜓)
& ⊢ (∃𝑥𝜑 → (𝜑 → 𝜓)) ⇒ ⊢ (∃𝑥𝜑 → 𝜓) |
|
Theorem | sb5ALT 39048* |
Equivalence for substitution. Alternate proof of sb5 2458.
This proof is
sb5ALTVD 39463 automatically translated and minimized.
(Contributed by Alan
Sare, 21-Apr-2013.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ ([𝑦 / 𝑥]𝜑 ↔ ∃𝑥(𝑥 = 𝑦 ∧ 𝜑)) |
|
Theorem | eexinst01 39049 |
exinst01 39167 without virtual deductions. (Contributed by
Alan Sare,
21-Apr-2013.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ ∃𝑥𝜓
& ⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → ∀𝑥𝜑)
& ⊢ (𝜒 → ∀𝑥𝜒) ⇒ ⊢ (𝜑 → 𝜒) |
|
Theorem | eexinst11 39050 |
exinst11 39168 without virtual deductions. (Contributed by
Alan Sare,
21-Apr-2013.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜑 → ∃𝑥𝜓)
& ⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜑 → ∀𝑥𝜑)
& ⊢ (𝜒 → ∀𝑥𝜒) ⇒ ⊢ (𝜑 → 𝜒) |
|
Theorem | vk15.4j 39051 |
Excercise 4j of Unit 15 of "Understanding Symbolic Logic", Fifth
Edition
(2008), by Virginia Klenk. This proof is the minimized Hilbert-style
axiomatic version of the Fitch-style Natural Deduction proof found on
page 442 of Klenk and was automatically derived from that proof.
vk15.4j 39051 is vk15.4jVD 39464 automatically translated and minimized.
(Contributed by Alan Sare, 21-Apr-2013.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ ¬
(∃𝑥 ¬ 𝜑 ∧ ∃𝑥(𝜓 ∧ ¬ 𝜒)) & ⊢ (∀𝑥𝜒 → ¬ ∃𝑥(𝜃 ∧ 𝜏)) & ⊢ ¬
∀𝑥(𝜏 → 𝜑) ⇒ ⊢ (¬ ∃𝑥 ¬ 𝜃 → ¬ ∀𝑥𝜓) |
|
Theorem | notnotrALT 39052 |
Converse of double negation. Alternate proof of notnotr 125. This proof
is notnotrALTVD 39465 automatically translated and minimized.
(Contributed by
Alan Sare, 21-Apr-2013.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (¬ ¬
𝜑 → 𝜑) |
|
Theorem | con3ALT2 39053 |
Contraposition. Alternate proof of con3 149. This proof is con3ALTVD 39466
automatically translated and minimized. (Contributed by Alan Sare,
21-Apr-2013.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ ((𝜑 → 𝜓) → (¬ 𝜓 → ¬ 𝜑)) |
|
Theorem | ssralv2 39054* |
Quantification restricted to a subclass for two quantifiers. ssralv 3699
for two quantifiers. The proof of ssralv2 39054 was automatically generated
by minimizing the automatically translated proof of ssralv2VD 39416. The
automatic translation is by the tools program
translatewithout_overwriting.cmd.
(Contributed by Alan Sare,
18-Feb-2012.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ ((𝐴 ⊆ 𝐵 ∧ 𝐶 ⊆ 𝐷) → (∀𝑥 ∈ 𝐵 ∀𝑦 ∈ 𝐷 𝜑 → ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐶 𝜑)) |
|
Theorem | sbc3or 39055 |
sbcor 3512 with a 3-disjuncts. This proof is sbc3orgVD 39400 automatically
translated and minimized. (Contributed by Alan Sare, 31-Dec-2011.)
(Revised by NM, 24-Aug-2018.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ ([𝐴 / 𝑥](𝜑 ∨ 𝜓 ∨ 𝜒) ↔ ([𝐴 / 𝑥]𝜑 ∨ [𝐴 / 𝑥]𝜓 ∨ [𝐴 / 𝑥]𝜒)) |
|
Theorem | sbcangOLD 39056 |
Distribution of class substitution over conjunction. (Contributed by
NM, 21-May-2004.) Obsolete as of 17-Aug-2018. Use sbcan 3511 instead.
(New usage is discouraged.) (Proof modification is discouraged.)
|
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥](𝜑 ∧ 𝜓) ↔ ([𝐴 / 𝑥]𝜑 ∧ [𝐴 / 𝑥]𝜓))) |
|
Theorem | sbcorgOLD 39057 |
Distribution of class substitution over disjunction. (Contributed by
NM, 21-May-2004.) Obsolete as of 17-Aug-2018. Use sbcor 3512 instead.
(New usage is discouraged.) (Proof modification is discouraged.)
|
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥](𝜑 ∨ 𝜓) ↔ ([𝐴 / 𝑥]𝜑 ∨ [𝐴 / 𝑥]𝜓))) |
|
Theorem | sbcbiiOLD 39058 |
Formula-building inference rule for class substitution. (Contributed by
NM, 11-Nov-2005.) Obsolete as of 17-Aug-2018. Use sbcbii 3524 instead.
(New usage is discouraged.) (Proof modification is discouraged.)
|
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥]𝜑 ↔ [𝐴 / 𝑥]𝜓)) |
|
Theorem | sbc3orgOLD 39059 |
sbcorgOLD 39057 with a 3-disjuncts. This proof is sbc3orgVD 39400 automatically
translated and minimized. (Contributed by Alan Sare, 31-Dec-2011.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥](𝜑 ∨ 𝜓 ∨ 𝜒) ↔ ([𝐴 / 𝑥]𝜑 ∨ [𝐴 / 𝑥]𝜓 ∨ [𝐴 / 𝑥]𝜒))) |
|
Theorem | alrim3con13v 39060* |
Closed form of alrimi 2120 with 2 additional conjuncts having no
occurrences of the quantifying variable. This proof is
19.21a3con13vVD 39401 automatically translated and minimized.
(Contributed
by Alan Sare, 31-Dec-2011.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ ((𝜑 → ∀𝑥𝜑) → ((𝜓 ∧ 𝜑 ∧ 𝜒) → ∀𝑥(𝜓 ∧ 𝜑 ∧ 𝜒))) |
|
Theorem | rspsbc2 39061* |
rspsbc 3551 with two quantifying variables. This proof
is rspsbc2VD 39404
automatically translated and minimized. (Contributed by Alan Sare,
31-Dec-2011.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝐴 ∈ 𝐵 → (𝐶 ∈ 𝐷 → (∀𝑥 ∈ 𝐵 ∀𝑦 ∈ 𝐷 𝜑 → [𝐶 / 𝑦][𝐴 / 𝑥]𝜑))) |
|
Theorem | sbcoreleleq 39062* |
Substitution of a setvar variable for another setvar variable in a
3-conjunct formula. Derived automatically from sbcoreleleqVD 39409.
(Contributed by Alan Sare, 31-Dec-2011.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑦](𝑥 ∈ 𝑦 ∨ 𝑦 ∈ 𝑥 ∨ 𝑥 = 𝑦) ↔ (𝑥 ∈ 𝐴 ∨ 𝐴 ∈ 𝑥 ∨ 𝑥 = 𝐴))) |
|
Theorem | tratrb 39063* |
If a class is transitive and any two distinct elements of the class are
E-comparable, then every element of that class is transitive. Derived
automatically from tratrbVD 39411. (Contributed by Alan Sare,
31-Dec-2011.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ ((Tr 𝐴 ∧ ∀𝑥 ∈ 𝐴 ∀𝑦 ∈ 𝐴 (𝑥 ∈ 𝑦 ∨ 𝑦 ∈ 𝑥 ∨ 𝑥 = 𝑦) ∧ 𝐵 ∈ 𝐴) → Tr 𝐵) |
|
Theorem | ordelordALT 39064 |
An element of an ordinal class is ordinal. Proposition 7.6 of
[TakeutiZaring] p. 36. This is
an alternate proof of ordelord 5783 using
the Axiom of Regularity indirectly through dford2 8555. dford2 is a weaker
definition of ordinal number. Given the Axiom of Regularity, it need
not be assumed that E Fr 𝐴 because this is inferred by the
Axiom of
Regularity. ordelordALT 39064 is ordelordALTVD 39417 without virtual deductions
and was automatically derived from ordelordALTVD 39417 using the tools
program translate..without..overwriting.cmd and Metamath's minimize
command. (Contributed by Alan Sare, 18-Feb-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ ((Ord 𝐴 ∧ 𝐵 ∈ 𝐴) → Ord 𝐵) |
|
Theorem | sbcim2g 39065 |
Distribution of class substitution over a left-nested implication.
Similar to sbcimg 3510. sbcim2g 39065 is sbcim2gVD 39425 without virtual deductions
and was automatically derived from sbcim2gVD 39425 using the tools program
translate..without..overwriting.cmd and Metamath's minimize command.
(Contributed by Alan Sare, 18-Mar-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥](𝜑 → (𝜓 → 𝜒)) ↔ ([𝐴 / 𝑥]𝜑 → ([𝐴 / 𝑥]𝜓 → [𝐴 / 𝑥]𝜒)))) |
|
Theorem | sbcbi 39066 |
Implication form of sbcbiiOLD 39058. sbcbi 39066 is sbcbiVD 39426 without virtual
deductions and was automatically derived from sbcbiVD 39426 using the tools
program translate..without..overwriting.cmd and Metamath's minimize
command. (Contributed by Alan Sare, 18-Mar-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (𝐴 ∈ 𝑉 → (∀𝑥(𝜑 ↔ 𝜓) → ([𝐴 / 𝑥]𝜑 ↔ [𝐴 / 𝑥]𝜓))) |
|
Theorem | trsbc 39067* |
Formula-building inference rule for class substitution, substituting a
class variable for the setvar variable of the transitivity predicate.
trsbc 39067 is trsbcVD 39427 without virtual deductions and was
automatically
derived from trsbcVD 39427 using the tools program
translate..without..overwriting.cmd and Metamath's minimize command.
(Contributed by Alan Sare, 18-Mar-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥]Tr 𝑥 ↔ Tr 𝐴)) |
|
Theorem | truniALT 39068* |
The union of a class of transitive sets is transitive. Alternate proof
of truni 4800. truniALT 39068 is truniALTVD 39428 without virtual deductions and
was automatically derived from truniALTVD 39428 using the tools program
translate..without..overwriting.cmd and Metamath's minimize command.
(Contributed by Alan Sare, 18-Mar-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (∀𝑥 ∈ 𝐴 Tr 𝑥 → Tr ∪
𝐴) |
|
Theorem | sbcalgOLD 39069* |
Move universal quantifier in and out of class substitution.
(Contributed by NM, 16-Jan-2004.) Obsolete as of 17-Aug-2018. Use
sbcal 3518 instead. (New usage is discouraged.)
(Proof modification is discouraged.)
|
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑦]∀𝑥𝜑 ↔ ∀𝑥[𝐴 / 𝑦]𝜑)) |
|
Theorem | sbcexgOLD 39070* |
Move existential quantifier in and out of class substitution.
(Contributed by NM, 21-May-2004.) Obsolete as of 17-Aug-2018. Use
sbcex2 3519 instead. (New usage is discouraged.)
(Proof modification is discouraged.)
|
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑦]∃𝑥𝜑 ↔ ∃𝑥[𝐴 / 𝑦]𝜑)) |
|
Theorem | sbcel12gOLD 39071 |
Distribute proper substitution through a membership relation.
(Contributed by NM, 10-Nov-2005.) (Proof shortened by Andrew Salmon,
29-Jun-2011.) Obsolete as of 18-Aug-2018. Use sbcel12 4016 instead.
(New usage is discouraged.) (Proof modification is discouraged.)
|
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥]𝐵 ∈ 𝐶 ↔ ⦋𝐴 / 𝑥⦌𝐵 ∈ ⦋𝐴 / 𝑥⦌𝐶)) |
|
Theorem | sbcel2gOLD 39072* |
Move proper substitution in and out of a membership relation.
(Contributed by NM, 14-Nov-2005.) Obsolete as of 18-Aug-2018. Use
sbcel2 4022 instead. (New usage is discouraged.)
(Proof modification is discouraged.)
|
⊢ (𝐴 ∈ 𝑉 → ([𝐴 / 𝑥]𝐵 ∈ 𝐶 ↔ 𝐵 ∈ ⦋𝐴 / 𝑥⦌𝐶)) |
|
Theorem | sbcssOLD 39073 |
Distribute proper substitution through a subclass relation. This
theorem was automatically derived from sbcssgVD 39433. (Contributed by
Alan Sare, 22-Jul-2012.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝐴 ∈ 𝐵 → ([𝐴 / 𝑥]𝐶 ⊆ 𝐷 ↔ ⦋𝐴 / 𝑥⦌𝐶 ⊆ ⦋𝐴 / 𝑥⦌𝐷)) |
|
Theorem | onfrALTlem5 39074* |
Lemma for onfrALT 39081. (Contributed by Alan Sare, 22-Jul-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢
([(𝑎 ∩
𝑥) / 𝑏]((𝑏 ⊆ (𝑎 ∩ 𝑥) ∧ 𝑏 ≠ ∅) → ∃𝑦 ∈ 𝑏 (𝑏 ∩ 𝑦) = ∅) ↔ (((𝑎 ∩ 𝑥) ⊆ (𝑎 ∩ 𝑥) ∧ (𝑎 ∩ 𝑥) ≠ ∅) → ∃𝑦 ∈ (𝑎 ∩ 𝑥)((𝑎 ∩ 𝑥) ∩ 𝑦) = ∅)) |
|
Theorem | onfrALTlem4 39075* |
Lemma for onfrALT 39081. (Contributed by Alan Sare, 22-Jul-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ ([𝑦 / 𝑥](𝑥 ∈ 𝑎 ∧ (𝑎 ∩ 𝑥) = ∅) ↔ (𝑦 ∈ 𝑎 ∧ (𝑎 ∩ 𝑦) = ∅)) |
|
Theorem | onfrALTlem3 39076* |
Lemma for onfrALT 39081. (Contributed by Alan Sare, 22-Jul-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ ((𝑎 ⊆ On ∧ 𝑎 ≠ ∅) → ((𝑥 ∈ 𝑎 ∧ ¬ (𝑎 ∩ 𝑥) = ∅) → ∃𝑦 ∈ (𝑎 ∩ 𝑥)((𝑎 ∩ 𝑥) ∩ 𝑦) = ∅)) |
|
Theorem | ggen31 39077* |
gen31 39163 without virtual deductions. (Contributed by
Alan Sare,
22-Jul-2012.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜑 → (𝜓 → (𝜒 → 𝜃))) ⇒ ⊢ (𝜑 → (𝜓 → (𝜒 → ∀𝑥𝜃))) |
|
Theorem | onfrALTlem2 39078* |
Lemma for onfrALT 39081. (Contributed by Alan Sare, 22-Jul-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ ((𝑎 ⊆ On ∧ 𝑎 ≠ ∅) → ((𝑥 ∈ 𝑎 ∧ ¬ (𝑎 ∩ 𝑥) = ∅) → ∃𝑦 ∈ 𝑎 (𝑎 ∩ 𝑦) = ∅)) |
|
Theorem | cbvexsv 39079* |
A theorem pertaining to the substitution for an existentially quantified
variable when the substituted variable does not occur in the quantified
wff. (Contributed by Alan Sare, 22-Jul-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (∃𝑥𝜑 ↔ ∃𝑦[𝑦 / 𝑥]𝜑) |
|
Theorem | onfrALTlem1 39080* |
Lemma for onfrALT 39081. (Contributed by Alan Sare, 22-Jul-2012.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ ((𝑎 ⊆ On ∧ 𝑎 ≠ ∅) → ((𝑥 ∈ 𝑎 ∧ (𝑎 ∩ 𝑥) = ∅) → ∃𝑦 ∈ 𝑎 (𝑎 ∩ 𝑦) = ∅)) |
|
Theorem | onfrALT 39081 |
The epsilon relation is foundational on the class of ordinal numbers.
onfrALT 39081 is an alternate proof of onfr 5801.
onfrALTVD 39441 is the Virtual
Deduction proof from which onfrALT 39081 is derived. The Virtual Deduction
proof mirrors the working proof of onfr 5801
which is the main part of the
proof of Theorem 7.12 of the first edition of TakeutiZaring. The proof
of the corresponding Proposition 7.12 of [TakeutiZaring] p. 38 (second
edition) does not contain the working proof equivalent of onfrALTVD 39441.
This theorem does not rely on the Axiom of Regularity. (Contributed by
Alan Sare, 22-Jul-2012.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ E Fr
On |
|
Theorem | csbeq2gOLD 39082 |
Formula-building implication rule for class substitution. Closed form of
csbeq2i 4026. csbeq2gOLD 39082 is derived from the virtual deduction proof
csbeq2gVD 39442. (Contributed by Alan Sare, 10-Nov-2012.)
Obsolete version
of csbeq2 3570 as of 11-Oct-2018. (Proof modification is
discouraged.)
(New usage is discouraged.)
|
⊢ (𝐴 ∈ 𝑉 → (∀𝑥 𝐵 = 𝐶 → ⦋𝐴 / 𝑥⦌𝐵 = ⦋𝐴 / 𝑥⦌𝐶)) |
|
Theorem | 19.41rg 39083 |
Closed form of right-to-left implication of 19.41 2141, Theorem 19.41 of
[Margaris] p. 90. Derived from 19.41rgVD 39452. (Contributed by Alan Sare,
8-Feb-2014.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (∀𝑥(𝜓 → ∀𝑥𝜓) → ((∃𝑥𝜑 ∧ 𝜓) → ∃𝑥(𝜑 ∧ 𝜓))) |
|
Theorem | opelopab4 39084* |
Ordered pair membership in a class abstraction of pairs. Compare to
elopab 5012. (Contributed by Alan Sare, 8-Feb-2014.)
(Revised by Mario
Carneiro, 6-May-2015.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (〈𝑢, 𝑣〉 ∈ {〈𝑥, 𝑦〉 ∣ 𝜑} ↔ ∃𝑥∃𝑦((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) ∧ 𝜑)) |
|
Theorem | 2pm13.193 39085 |
pm13.193 38929 for two variables. pm13.193 38929 is Theorem *13.193 in
[WhiteheadRussell] p. 179.
Derived from 2pm13.193VD 39453. (Contributed by
Alan Sare, 8-Feb-2014.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) ∧ [𝑢 / 𝑥][𝑣 / 𝑦]𝜑) ↔ ((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) ∧ 𝜑)) |
|
Theorem | hbntal 39086 |
A closed form of hbn 2184. hbnt 2182 is another closed form of hbn 2184.
(Contributed by Alan Sare, 8-Feb-2014.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (∀𝑥(𝜑 → ∀𝑥𝜑) → ∀𝑥(¬ 𝜑 → ∀𝑥 ¬ 𝜑)) |
|
Theorem | hbimpg 39087 |
A closed form of hbim 2165. Derived from hbimpgVD 39454. (Contributed by Alan
Sare, 8-Feb-2014.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢
((∀𝑥(𝜑 → ∀𝑥𝜑) ∧ ∀𝑥(𝜓 → ∀𝑥𝜓)) → ∀𝑥((𝜑 → 𝜓) → ∀𝑥(𝜑 → 𝜓))) |
|
Theorem | hbalg 39088 |
Closed form of hbal 2076. Derived from hbalgVD 39455. (Contributed by Alan
Sare, 8-Feb-2014.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (∀𝑦(𝜑 → ∀𝑥𝜑) → ∀𝑦(∀𝑦𝜑 → ∀𝑥∀𝑦𝜑)) |
|
Theorem | hbexg 39089 |
Closed form of nfex 2192. Derived from hbexgVD 39456. (Contributed by Alan
Sare, 8-Feb-2014.) (Revised by Mario Carneiro, 12-Dec-2016.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (∀𝑥∀𝑦(𝜑 → ∀𝑥𝜑) → ∀𝑥∀𝑦(∃𝑦𝜑 → ∀𝑥∃𝑦𝜑)) |
|
Theorem | ax6e2eq 39090* |
Alternate form of ax6e 2286 for non-distinct 𝑥, 𝑦 and 𝑢 = 𝑣.
ax6e2eq 39090 is derived from ax6e2eqVD 39457. (Contributed by Alan Sare,
25-Mar-2014.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (∀𝑥 𝑥 = 𝑦 → (𝑢 = 𝑣 → ∃𝑥∃𝑦(𝑥 = 𝑢 ∧ 𝑦 = 𝑣))) |
|
Theorem | ax6e2nd 39091* |
If at least two sets exist (dtru 4887) , then the same is true expressed
in an alternate form similar to the form of ax6e 2286.
ax6e2nd 39091 is
derived from ax6e2ndVD 39458. (Contributed by Alan Sare, 25-Mar-2014.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (¬
∀𝑥 𝑥 = 𝑦 → ∃𝑥∃𝑦(𝑥 = 𝑢 ∧ 𝑦 = 𝑣)) |
|
Theorem | ax6e2ndeq 39092* |
"At least two sets exist" expressed in the form of dtru 4887
is logically
equivalent to the same expressed in a form similar to ax6e 2286
if dtru 4887
is false implies 𝑢 = 𝑣. ax6e2ndeq 39092 is derived from ax6e2ndeqVD 39459.
(Contributed by Alan Sare, 25-Mar-2014.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ ((¬
∀𝑥 𝑥 = 𝑦 ∨ 𝑢 = 𝑣) ↔ ∃𝑥∃𝑦(𝑥 = 𝑢 ∧ 𝑦 = 𝑣)) |
|
Theorem | 2sb5nd 39093* |
Equivalence for double substitution 2sb5 2471 without distinct 𝑥,
𝑦 requirement. 2sb5nd 39093 is derived from 2sb5ndVD 39460. (Contributed
by Alan Sare, 30-Apr-2014.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ ((¬
∀𝑥 𝑥 = 𝑦 ∨ 𝑢 = 𝑣) → ([𝑢 / 𝑥][𝑣 / 𝑦]𝜑 ↔ ∃𝑥∃𝑦((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) ∧ 𝜑))) |
|
Theorem | 2uasbanh 39094* |
Distribute the unabbreviated form of proper substitution in and out of a
conjunction. 2uasbanh 39094 is derived from 2uasbanhVD 39461. (Contributed by
Alan Sare, 31-May-2014.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜒 ↔ (∃𝑥∃𝑦((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) ∧ 𝜑) ∧ ∃𝑥∃𝑦((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) ∧ 𝜓))) ⇒ ⊢ (∃𝑥∃𝑦((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) ∧ (𝜑 ∧ 𝜓)) ↔ (∃𝑥∃𝑦((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) ∧ 𝜑) ∧ ∃𝑥∃𝑦((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) ∧ 𝜓))) |
|
Theorem | 2uasban 39095* |
Distribute the unabbreviated form of proper substitution in and out of a
conjunction. (Contributed by Alan Sare, 31-May-2014.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (∃𝑥∃𝑦((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) ∧ (𝜑 ∧ 𝜓)) ↔ (∃𝑥∃𝑦((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) ∧ 𝜑) ∧ ∃𝑥∃𝑦((𝑥 = 𝑢 ∧ 𝑦 = 𝑣) ∧ 𝜓))) |
|
Theorem | e2ebind 39096 |
Absorption of an existential quantifier of a double existential quantifier
of non-distinct variables. e2ebind 39096 is derived from e2ebindVD 39462.
(Contributed by Alan Sare, 27-Nov-2014.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (∀𝑥 𝑥 = 𝑦 → (∃𝑥∃𝑦𝜑 ↔ ∃𝑦𝜑)) |
|
Theorem | elpwgded 39097 |
elpwgdedVD 39467 in conventional notation. (Contributed by
Alan Sare,
23-Apr-2015.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜑 → 𝐴 ∈ V) & ⊢ (𝜓 → 𝐴 ⊆ 𝐵) ⇒ ⊢ ((𝜑 ∧ 𝜓) → 𝐴 ∈ 𝒫 𝐵) |
|
Theorem | trelded 39098 |
Deduction form of trel 4792. In a transitive class, the membership
relation is transitive. (Contributed by Alan Sare, 3-Dec-2015.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (𝜑 → Tr 𝐴)
& ⊢ (𝜓 → 𝐵 ∈ 𝐶)
& ⊢ (𝜒 → 𝐶 ∈ 𝐴) ⇒ ⊢ ((𝜑 ∧ 𝜓 ∧ 𝜒) → 𝐵 ∈ 𝐴) |
|
Theorem | jaoded 39099 |
Deduction form of jao 533. Disjunction of antecedents. (Contributed by
Alan Sare, 3-Dec-2015.) (Proof modification is discouraged.)
(New usage is discouraged.)
|
⊢ (𝜑 → (𝜓 → 𝜒)) & ⊢ (𝜃 → (𝜏 → 𝜒)) & ⊢ (𝜂 → (𝜓 ∨ 𝜏)) ⇒ ⊢ ((𝜑 ∧ 𝜃 ∧ 𝜂) → 𝜒) |
|
Theorem | sbtT 39100 |
A substitution into a theorem remains true. sbt 2447
with the existence of
no virtual hypotheses for the hypothesis expressed as the empty virtual
hypothesis collection. (Contributed by Alan Sare, 4-Feb-2017.)
(Proof modification is discouraged.) (New usage is discouraged.)
|
⊢ (⊤
→ 𝜑) ⇒ ⊢ [𝑦 / 𝑥]𝜑 |