![]() |
Metamath
Proof Explorer Theorem List (p. 27 of 431) | < 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-28105) |
![]() (28106-29630) |
![]() (29631-43082) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | sbalv 2601* | Quantify with new variable inside substitution. (Contributed by NM, 18-Aug-1993.) |
⊢ ([𝑦 / 𝑥]𝜑 ↔ 𝜓) ⇒ ⊢ ([𝑦 / 𝑥]∀𝑧𝜑 ↔ ∀𝑧𝜓) | ||
Theorem | sbco4lem 2602* | Lemma for sbco4 2603. It replaces the temporary variable 𝑣 with another temporary variable 𝑤. (Contributed by Jim Kingdon, 26-Sep-2018.) |
⊢ ([𝑥 / 𝑣][𝑦 / 𝑥][𝑣 / 𝑦]𝜑 ↔ [𝑥 / 𝑤][𝑦 / 𝑥][𝑤 / 𝑦]𝜑) | ||
Theorem | sbco4 2603* | Two ways of exchanging two variables. Both sides of the biconditional exchange 𝑥 and 𝑦, either via two temporary variables 𝑢 and 𝑣, or a single temporary 𝑤. (Contributed by Jim Kingdon, 25-Sep-2018.) |
⊢ ([𝑦 / 𝑢][𝑥 / 𝑣][𝑢 / 𝑥][𝑣 / 𝑦]𝜑 ↔ [𝑥 / 𝑤][𝑦 / 𝑥][𝑤 / 𝑦]𝜑) | ||
Theorem | 2sb8e 2604* | An equivalent expression for double existence. (Contributed by Wolf Lammen, 2-Nov-2019.) |
⊢ (∃𝑥∃𝑦𝜑 ↔ ∃𝑧∃𝑤[𝑧 / 𝑥][𝑤 / 𝑦]𝜑) | ||
Theorem | exsb 2605* | An equivalent expression for existence. (Contributed by NM, 2-Feb-2005.) |
⊢ (∃𝑥𝜑 ↔ ∃𝑦∀𝑥(𝑥 = 𝑦 → 𝜑)) | ||
Theorem | 2exsb 2606* | An equivalent expression for double existence. (Contributed by NM, 2-Feb-2005.) (Proof shortened by Wolf Lammen, 30-Sep-2018.) |
⊢ (∃𝑥∃𝑦𝜑 ↔ ∃𝑧∃𝑤∀𝑥∀𝑦((𝑥 = 𝑧 ∧ 𝑦 = 𝑤) → 𝜑)) | ||
Syntax | weu 2607 | Extend wff definition to include existential uniqueness ("there exists a unique 𝑥 such that 𝜑"). |
wff ∃!𝑥𝜑 | ||
Syntax | wmo 2608 | Extend wff definition to include uniqueness ("there exists at most one 𝑥 such that 𝜑"). |
wff ∃*𝑥𝜑 | ||
Theorem | eujust 2609* | A soundness justification theorem for df-eu 2611, showing that the definition is equivalent to itself with its dummy variable renamed. Note that 𝑦 and 𝑧 needn't be distinct variables. See eujustALT 2610 for a proof that provides an example of how it can be achieved through the use of dvelim 2477. (Contributed by NM, 11-Mar-2010.) (Proof shortened by Andrew Salmon, 9-Jul-2011.) |
⊢ (∃𝑦∀𝑥(𝜑 ↔ 𝑥 = 𝑦) ↔ ∃𝑧∀𝑥(𝜑 ↔ 𝑥 = 𝑧)) | ||
Theorem | eujustALT 2610* | Alternate proof of eujust 2609 illustrating the use of dvelim 2477. (Contributed by NM, 11-Mar-2010.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (∃𝑦∀𝑥(𝜑 ↔ 𝑥 = 𝑦) ↔ ∃𝑧∀𝑥(𝜑 ↔ 𝑥 = 𝑧)) | ||
Definition | df-eu 2611* | Define existential uniqueness, i.e. "there exists exactly one 𝑥 such that 𝜑." Definition 10.1 of [BellMachover] p. 97; also Definition *14.02 of [WhiteheadRussell] p. 175. Other possible definitions are given by eu1 2648, eu2 2647, eu3v 2635, and eu5 2633 (which in some cases we show with a hypothesis 𝜑 → ∀𝑦𝜑 in place of a distinct variable condition on 𝑦 and 𝜑). Double uniqueness is tricky: ∃!𝑥∃!𝑦𝜑 does not mean "exactly one 𝑥 and one 𝑦 " (see 2eu4 2694). (Contributed by NM, 12-Aug-1993.) |
⊢ (∃!𝑥𝜑 ↔ ∃𝑦∀𝑥(𝜑 ↔ 𝑥 = 𝑦)) | ||
Definition | df-mo 2612 | Define "there exists at most one 𝑥 such that 𝜑." Here we define it in terms of existential uniqueness. Notation of [BellMachover] p. 460, whose definition we show as mo3 2645. For other possible definitions see mo2 2616 and mo4 2655. (Contributed by NM, 8-Mar-1995.) |
⊢ (∃*𝑥𝜑 ↔ (∃𝑥𝜑 → ∃!𝑥𝜑)) | ||
Theorem | euequ1 2613* | Equality has existential uniqueness. Special case of eueq1 3520 proved using only predicate calculus. The proof needs 𝑦 = 𝑧 be free of 𝑥. This is ensured by having 𝑥 and 𝑦 be distinct. Alternately, a distinctor ¬ ∀𝑥𝑥 = 𝑦 could have been used instead. (Contributed by Stefan Allan, 4-Dec-2008.) (Proof shortened by Wolf Lammen, 8-Sep-2019.) |
⊢ ∃!𝑥 𝑥 = 𝑦 | ||
Theorem | mo2v 2614* | Alternate definition of "at most one." Unlike mo2 2616, which is slightly more general, it does not depend on ax-11 2183 and ax-13 2391, whence it is preferable within predicate logic. Elsewhere, most theorems depend on these axioms anyway, so this advantage is no longer important. (Contributed by Wolf Lammen, 27-May-2019.) (Proof shortened by Wolf Lammen, 10-Nov-2019.) |
⊢ (∃*𝑥𝜑 ↔ ∃𝑦∀𝑥(𝜑 → 𝑥 = 𝑦)) | ||
Theorem | euf 2615* | A version of the existential uniqueness definition with a hypothesis instead of a distinct variable condition. (Contributed by NM, 12-Aug-1993.) (Proof shortened by Wolf Lammen, 30-Oct-2018.) |
⊢ Ⅎ𝑦𝜑 ⇒ ⊢ (∃!𝑥𝜑 ↔ ∃𝑦∀𝑥(𝜑 ↔ 𝑥 = 𝑦)) | ||
Theorem | mo2 2616* | Alternate definition of "at most one." (Contributed by NM, 8-Mar-1995.) Restrict dummy variable z. (Revised by Wolf Lammen, 28-May-2019.) |
⊢ Ⅎ𝑦𝜑 ⇒ ⊢ (∃*𝑥𝜑 ↔ ∃𝑦∀𝑥(𝜑 → 𝑥 = 𝑦)) | ||
Theorem | nfeu1 2617 | Bound-variable hypothesis builder for uniqueness. (Contributed by NM, 9-Jul-1994.) (Revised by Mario Carneiro, 7-Oct-2016.) |
⊢ Ⅎ𝑥∃!𝑥𝜑 | ||
Theorem | nfmo1 2618 | Bound-variable hypothesis builder for "at most one." (Contributed by NM, 8-Mar-1995.) (Revised by Mario Carneiro, 7-Oct-2016.) |
⊢ Ⅎ𝑥∃*𝑥𝜑 | ||
Theorem | nfeud2 2619 | Bound-variable hypothesis builder for uniqueness. (Contributed by Mario Carneiro, 14-Nov-2016.) (Proof shortened by Wolf Lammen, 4-Oct-2018.) |
⊢ Ⅎ𝑦𝜑 & ⊢ ((𝜑 ∧ ¬ ∀𝑥 𝑥 = 𝑦) → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → Ⅎ𝑥∃!𝑦𝜓) | ||
Theorem | nfmod2 2620 | Bound-variable hypothesis builder for "at most one." (Contributed by Mario Carneiro, 14-Nov-2016.) |
⊢ Ⅎ𝑦𝜑 & ⊢ ((𝜑 ∧ ¬ ∀𝑥 𝑥 = 𝑦) → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → Ⅎ𝑥∃*𝑦𝜓) | ||
Theorem | nfeud 2621 | Deduction version of nfeu 2623. (Contributed by NM, 15-Feb-2013.) (Revised by Mario Carneiro, 7-Oct-2016.) |
⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → Ⅎ𝑥∃!𝑦𝜓) | ||
Theorem | nfmod 2622 | Bound-variable hypothesis builder for "at most one." (Contributed by Mario Carneiro, 14-Nov-2016.) |
⊢ Ⅎ𝑦𝜑 & ⊢ (𝜑 → Ⅎ𝑥𝜓) ⇒ ⊢ (𝜑 → Ⅎ𝑥∃*𝑦𝜓) | ||
Theorem | nfeu 2623 | Bound-variable hypothesis builder for uniqueness. Note that 𝑥 and 𝑦 needn't be distinct. (Contributed by NM, 8-Mar-1995.) (Revised by Mario Carneiro, 7-Oct-2016.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ Ⅎ𝑥∃!𝑦𝜑 | ||
Theorem | nfmo 2624 | Bound-variable hypothesis builder for "at most one." (Contributed by NM, 9-Mar-1995.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ Ⅎ𝑥∃*𝑦𝜑 | ||
Theorem | eubid 2625 | Formula-building rule for uniqueness quantifier (deduction rule). (Contributed by NM, 9-Jul-1994.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (∃!𝑥𝜓 ↔ ∃!𝑥𝜒)) | ||
Theorem | mobid 2626 | Formula-building rule for "at most one" quantifier (deduction rule). (Contributed by NM, 8-Mar-1995.) |
⊢ Ⅎ𝑥𝜑 & ⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (∃*𝑥𝜓 ↔ ∃*𝑥𝜒)) | ||
Theorem | eubidv 2627* | Formula-building rule for uniqueness quantifier (deduction rule). (Contributed by NM, 9-Jul-1994.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (∃!𝑥𝜓 ↔ ∃!𝑥𝜒)) | ||
Theorem | mobidv 2628* | Formula-building rule for "at most one" quantifier (deduction rule). (Contributed by Mario Carneiro, 7-Oct-2016.) |
⊢ (𝜑 → (𝜓 ↔ 𝜒)) ⇒ ⊢ (𝜑 → (∃*𝑥𝜓 ↔ ∃*𝑥𝜒)) | ||
Theorem | eubii 2629 | Introduce uniqueness quantifier to both sides of an equivalence. (Contributed by NM, 9-Jul-1994.) (Revised by Mario Carneiro, 6-Oct-2016.) |
⊢ (𝜑 ↔ 𝜓) ⇒ ⊢ (∃!𝑥𝜑 ↔ ∃!𝑥𝜓) | ||
Theorem | mobii 2630 | Formula-building rule for "at most one" quantifier (inference rule). (Contributed by NM, 9-Mar-1995.) (Revised by Mario Carneiro, 17-Oct-2016.) |
⊢ (𝜓 ↔ 𝜒) ⇒ ⊢ (∃*𝑥𝜓 ↔ ∃*𝑥𝜒) | ||
Theorem | euex 2631 | Existential uniqueness implies existence. For a shorter proof using more axioms, see euexALT 2649. (Contributed by NM, 15-Sep-1993.) (Proof shortened by Andrew Salmon, 9-Jul-2011.) (Proof shortened by Wolf Lammen, 4-Dec-2018.) |
⊢ (∃!𝑥𝜑 → ∃𝑥𝜑) | ||
Theorem | exmo 2632 | Something exists or at most one exists. (Contributed by NM, 8-Mar-1995.) |
⊢ (∃𝑥𝜑 ∨ ∃*𝑥𝜑) | ||
Theorem | eu5 2633 | Uniqueness in terms of "at most one." Revised to reduce dependencies on axioms. (Contributed by NM, 23-Mar-1995.) (Proof shortened by Wolf Lammen, 25-May-2019.) |
⊢ (∃!𝑥𝜑 ↔ (∃𝑥𝜑 ∧ ∃*𝑥𝜑)) | ||
Theorem | exmoeu2 2634 | Existence implies "at most one" is equivalent to uniqueness. (Contributed by NM, 5-Apr-2004.) |
⊢ (∃𝑥𝜑 → (∃*𝑥𝜑 ↔ ∃!𝑥𝜑)) | ||
Theorem | eu3v 2635* | An alternate way to express existential uniqueness. (Contributed by NM, 8-Jul-1994.) Add a distinct variable condition on 𝜑. (Revised by Wolf Lammen, 29-May-2019.) |
⊢ (∃!𝑥𝜑 ↔ (∃𝑥𝜑 ∧ ∃𝑦∀𝑥(𝜑 → 𝑥 = 𝑦))) | ||
Theorem | eumo 2636 | Existential uniqueness implies "at most one." (Contributed by NM, 23-Mar-1995.) Reduce axiom usage. (Revised by GL, 19-Feb-2022.) |
⊢ (∃!𝑥𝜑 → ∃*𝑥𝜑) | ||
Theorem | eumoOLD 2637 | Obsolete proof of eumo 2636 as of 19-Feb-2022. (Contributed by NM, 23-Mar-1995.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (∃!𝑥𝜑 → ∃*𝑥𝜑) | ||
Theorem | eumoi 2638 | "At most one" inferred from existential uniqueness. (Contributed by NM, 5-Apr-1995.) |
⊢ ∃!𝑥𝜑 ⇒ ⊢ ∃*𝑥𝜑 | ||
Theorem | moabs 2639 | Absorption of existence condition by "at most one." (Contributed by NM, 4-Nov-2002.) |
⊢ (∃*𝑥𝜑 ↔ (∃𝑥𝜑 → ∃*𝑥𝜑)) | ||
Theorem | exmoeu 2640 | Existence in terms of "at most one" and uniqueness. (Contributed by NM, 5-Apr-2004.) (Proof shortened by Wolf Lammen, 5-Dec-2018.) |
⊢ (∃𝑥𝜑 ↔ (∃*𝑥𝜑 → ∃!𝑥𝜑)) | ||
Theorem | sb8eu 2641 | Variable substitution in uniqueness quantifier. (Contributed by NM, 7-Aug-1994.) (Revised by Mario Carneiro, 7-Oct-2016.) (Proof shortened by Wolf Lammen, 24-Aug-2019.) |
⊢ Ⅎ𝑦𝜑 ⇒ ⊢ (∃!𝑥𝜑 ↔ ∃!𝑦[𝑦 / 𝑥]𝜑) | ||
Theorem | sb8mo 2642 | Variable substitution for "at most one." (Contributed by Alexander van der Vekens, 17-Jun-2017.) |
⊢ Ⅎ𝑦𝜑 ⇒ ⊢ (∃*𝑥𝜑 ↔ ∃*𝑦[𝑦 / 𝑥]𝜑) | ||
Theorem | cbveu 2643 | Rule used to change bound variables, using implicit substitution. (Contributed by NM, 25-Nov-1994.) (Revised by Mario Carneiro, 7-Oct-2016.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃!𝑥𝜑 ↔ ∃!𝑦𝜓) | ||
Theorem | cbvmo 2644 | Rule used to change bound variables, using implicit substitution. (Contributed by NM, 9-Mar-1995.) (Revised by Andrew Salmon, 8-Jun-2011.) |
⊢ Ⅎ𝑦𝜑 & ⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃*𝑥𝜑 ↔ ∃*𝑦𝜓) | ||
Theorem | mo3 2645* | Alternate definition of "at most one." Definition of [BellMachover] p. 460, except that definition has the side condition that 𝑦 not occur in 𝜑 in place of our hypothesis. (Contributed by NM, 8-Mar-1995.) (Proof shortened by Wolf Lammen, 18-Aug-2019.) |
⊢ Ⅎ𝑦𝜑 ⇒ ⊢ (∃*𝑥𝜑 ↔ ∀𝑥∀𝑦((𝜑 ∧ [𝑦 / 𝑥]𝜑) → 𝑥 = 𝑦)) | ||
Theorem | mo 2646* | Equivalent definitions of "there exists at most one." (Contributed by NM, 7-Aug-1994.) (Revised by Mario Carneiro, 7-Oct-2016.) (Proof shortened by Wolf Lammen, 2-Dec-2018.) |
⊢ Ⅎ𝑦𝜑 ⇒ ⊢ (∃𝑦∀𝑥(𝜑 → 𝑥 = 𝑦) ↔ ∀𝑥∀𝑦((𝜑 ∧ [𝑦 / 𝑥]𝜑) → 𝑥 = 𝑦)) | ||
Theorem | eu2 2647* | An alternate way of defining existential uniqueness. Definition 6.10 of [TakeutiZaring] p. 26. (Contributed by NM, 8-Jul-1994.) (Proof shortened by Wolf Lammen, 2-Dec-2018.) |
⊢ Ⅎ𝑦𝜑 ⇒ ⊢ (∃!𝑥𝜑 ↔ (∃𝑥𝜑 ∧ ∀𝑥∀𝑦((𝜑 ∧ [𝑦 / 𝑥]𝜑) → 𝑥 = 𝑦))) | ||
Theorem | eu1 2648* | An alternate way to express uniqueness used by some authors. Exercise 2(b) of [Margaris] p. 110. (Contributed by NM, 20-Aug-1993.) (Revised by Mario Carneiro, 7-Oct-2016.) (Proof shortened by Wolf Lammen, 29-Oct-2018.) |
⊢ Ⅎ𝑦𝜑 ⇒ ⊢ (∃!𝑥𝜑 ↔ ∃𝑥(𝜑 ∧ ∀𝑦([𝑦 / 𝑥]𝜑 → 𝑥 = 𝑦))) | ||
Theorem | euexALT 2649 | Alternate proof of euex 2631. Shorter but uses more axioms. (Contributed by NM, 15-Sep-1993.) (Proof shortened by Andrew Salmon, 9-Jul-2011.) (Proof modification is discouraged.) (New usage is discouraged.) |
⊢ (∃!𝑥𝜑 → ∃𝑥𝜑) | ||
Theorem | euor 2650 | Introduce a disjunct into a uniqueness quantifier. (Contributed by NM, 21-Oct-2005.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ ((¬ 𝜑 ∧ ∃!𝑥𝜓) → ∃!𝑥(𝜑 ∨ 𝜓)) | ||
Theorem | euorv 2651* | Introduce a disjunct into a uniqueness quantifier. (Contributed by NM, 23-Mar-1995.) |
⊢ ((¬ 𝜑 ∧ ∃!𝑥𝜓) → ∃!𝑥(𝜑 ∨ 𝜓)) | ||
Theorem | euor2 2652 | Introduce or eliminate a disjunct in a uniqueness quantifier. (Contributed by NM, 21-Oct-2005.) (Proof shortened by Andrew Salmon, 9-Jul-2011.) (Proof shortened by Wolf Lammen, 27-Dec-2018.) |
⊢ (¬ ∃𝑥𝜑 → (∃!𝑥(𝜑 ∨ 𝜓) ↔ ∃!𝑥𝜓)) | ||
Theorem | sbmo 2653* | Substitution into "at most one". (Contributed by Jeff Madsen, 2-Sep-2009.) |
⊢ ([𝑦 / 𝑥]∃*𝑧𝜑 ↔ ∃*𝑧[𝑦 / 𝑥]𝜑) | ||
Theorem | mo4f 2654* | "At most one" expressed using implicit substitution. (Contributed by NM, 10-Apr-2004.) |
⊢ Ⅎ𝑥𝜓 & ⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃*𝑥𝜑 ↔ ∀𝑥∀𝑦((𝜑 ∧ 𝜓) → 𝑥 = 𝑦)) | ||
Theorem | mo4 2655* | "At most one" expressed using implicit substitution. (Contributed by NM, 26-Jul-1995.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃*𝑥𝜑 ↔ ∀𝑥∀𝑦((𝜑 ∧ 𝜓) → 𝑥 = 𝑦)) | ||
Theorem | eu4 2656* | Uniqueness using implicit substitution. (Contributed by NM, 26-Jul-1995.) |
⊢ (𝑥 = 𝑦 → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃!𝑥𝜑 ↔ (∃𝑥𝜑 ∧ ∀𝑥∀𝑦((𝜑 ∧ 𝜓) → 𝑥 = 𝑦))) | ||
Theorem | moim 2657 | "At most one" reverses implication. (Contributed by NM, 22-Apr-1995.) |
⊢ (∀𝑥(𝜑 → 𝜓) → (∃*𝑥𝜓 → ∃*𝑥𝜑)) | ||
Theorem | moimi 2658 | "At most one" reverses implication. (Contributed by NM, 15-Feb-2006.) |
⊢ (𝜑 → 𝜓) ⇒ ⊢ (∃*𝑥𝜓 → ∃*𝑥𝜑) | ||
Theorem | moa1 2659 | If an implication holds for at most one value, then its consequent holds for at most one value. See also ala1 1890 and exa1 1914. (Contributed by NM, 28-Jul-1995.) (Proof shortened by Wolf Lammen, 22-Dec-2018.) (Revised by BJ, 29-Mar-2021.) |
⊢ (∃*𝑥(𝜑 → 𝜓) → ∃*𝑥𝜓) | ||
Theorem | euimmo 2660 | Uniqueness implies "at most one" through reverse implication. (Contributed by NM, 22-Apr-1995.) |
⊢ (∀𝑥(𝜑 → 𝜓) → (∃!𝑥𝜓 → ∃*𝑥𝜑)) | ||
Theorem | euim 2661 | Add existential uniqueness quantifiers to an implication. Note the reversed implication in the antecedent. (Contributed by NM, 19-Oct-2005.) (Proof shortened by Andrew Salmon, 14-Jun-2011.) |
⊢ ((∃𝑥𝜑 ∧ ∀𝑥(𝜑 → 𝜓)) → (∃!𝑥𝜓 → ∃!𝑥𝜑)) | ||
Theorem | moan 2662 | "At most one" is still the case when a conjunct is added. (Contributed by NM, 22-Apr-1995.) |
⊢ (∃*𝑥𝜑 → ∃*𝑥(𝜓 ∧ 𝜑)) | ||
Theorem | moani 2663 | "At most one" is still true when a conjunct is added. (Contributed by NM, 9-Mar-1995.) |
⊢ ∃*𝑥𝜑 ⇒ ⊢ ∃*𝑥(𝜓 ∧ 𝜑) | ||
Theorem | moor 2664 | "At most one" is still the case when a disjunct is removed. (Contributed by NM, 5-Apr-2004.) |
⊢ (∃*𝑥(𝜑 ∨ 𝜓) → ∃*𝑥𝜑) | ||
Theorem | mooran1 2665 | "At most one" imports disjunction to conjunction. (Contributed by NM, 5-Apr-2004.) (Proof shortened by Andrew Salmon, 9-Jul-2011.) |
⊢ ((∃*𝑥𝜑 ∨ ∃*𝑥𝜓) → ∃*𝑥(𝜑 ∧ 𝜓)) | ||
Theorem | mooran2 2666 | "At most one" exports disjunction to conjunction. (Contributed by NM, 5-Apr-2004.) (Proof shortened by Andrew Salmon, 9-Jul-2011.) |
⊢ (∃*𝑥(𝜑 ∨ 𝜓) → (∃*𝑥𝜑 ∧ ∃*𝑥𝜓)) | ||
Theorem | moanim 2667 | Introduction of a conjunct into "at most one" quantifier. (Contributed by NM, 3-Dec-2001.) (Proof shortened by Wolf Lammen, 24-Dec-2018.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (∃*𝑥(𝜑 ∧ 𝜓) ↔ (𝜑 → ∃*𝑥𝜓)) | ||
Theorem | euan 2668 | Introduction of a conjunct into uniqueness quantifier. (Contributed by NM, 19-Feb-2005.) (Proof shortened by Andrew Salmon, 9-Jul-2011.) (Proof shortened by Wolf Lammen, 24-Dec-2018.) |
⊢ Ⅎ𝑥𝜑 ⇒ ⊢ (∃!𝑥(𝜑 ∧ 𝜓) ↔ (𝜑 ∧ ∃!𝑥𝜓)) | ||
Theorem | moanimv 2669* | Introduction of a conjunct into "at most one" quantifier. (Contributed by NM, 23-Mar-1995.) |
⊢ (∃*𝑥(𝜑 ∧ 𝜓) ↔ (𝜑 → ∃*𝑥𝜓)) | ||
Theorem | moanmo 2670 | Nested "at most one" quantifiers. (Contributed by NM, 25-Jan-2006.) |
⊢ ∃*𝑥(𝜑 ∧ ∃*𝑥𝜑) | ||
Theorem | moaneu 2671 | Nested "at most one" and uniqueness quantifiers. (Contributed by NM, 25-Jan-2006.) (Proof shortened by Wolf Lammen, 27-Dec-2018.) |
⊢ ∃*𝑥(𝜑 ∧ ∃!𝑥𝜑) | ||
Theorem | euanv 2672* | Introduction of a conjunct into uniqueness quantifier. (Contributed by NM, 23-Mar-1995.) |
⊢ (∃!𝑥(𝜑 ∧ 𝜓) ↔ (𝜑 ∧ ∃!𝑥𝜓)) | ||
Theorem | mopick 2673 | "At most one" picks a variable value, eliminating an existential quantifier. (Contributed by NM, 27-Jan-1997.) (Proof shortened by Wolf Lammen, 17-Sep-2019.) |
⊢ ((∃*𝑥𝜑 ∧ ∃𝑥(𝜑 ∧ 𝜓)) → (𝜑 → 𝜓)) | ||
Theorem | eupick 2674 | Existential uniqueness "picks" a variable value for which another wff is true. If there is only one thing 𝑥 such that 𝜑 is true, and there is also an 𝑥 (actually the same one) such that 𝜑 and 𝜓 are both true, then 𝜑 implies 𝜓 regardless of 𝑥. This theorem can be useful for eliminating existential quantifiers in a hypothesis. Compare Theorem *14.26 in [WhiteheadRussell] p. 192. (Contributed by NM, 10-Jul-1994.) |
⊢ ((∃!𝑥𝜑 ∧ ∃𝑥(𝜑 ∧ 𝜓)) → (𝜑 → 𝜓)) | ||
Theorem | eupicka 2675 | Version of eupick 2674 with closed formulas. (Contributed by NM, 6-Sep-2008.) |
⊢ ((∃!𝑥𝜑 ∧ ∃𝑥(𝜑 ∧ 𝜓)) → ∀𝑥(𝜑 → 𝜓)) | ||
Theorem | eupickb 2676 | Existential uniqueness "pick" showing wff equivalence. (Contributed by NM, 25-Nov-1994.) (Proof shortened by Wolf Lammen, 27-Dec-2018.) |
⊢ ((∃!𝑥𝜑 ∧ ∃!𝑥𝜓 ∧ ∃𝑥(𝜑 ∧ 𝜓)) → (𝜑 ↔ 𝜓)) | ||
Theorem | eupickbi 2677 | Theorem *14.26 in [WhiteheadRussell] p. 192. (Contributed by Andrew Salmon, 11-Jul-2011.) (Proof shortened by Wolf Lammen, 27-Dec-2018.) |
⊢ (∃!𝑥𝜑 → (∃𝑥(𝜑 ∧ 𝜓) ↔ ∀𝑥(𝜑 → 𝜓))) | ||
Theorem | mopick2 2678 | "At most one" can show the existence of a common value. In this case we can infer existence of conjunction from a conjunction of existence, and it is one way to achieve the converse of 19.40 1946. (Contributed by NM, 5-Apr-2004.) (Proof shortened by Andrew Salmon, 9-Jul-2011.) |
⊢ ((∃*𝑥𝜑 ∧ ∃𝑥(𝜑 ∧ 𝜓) ∧ ∃𝑥(𝜑 ∧ 𝜒)) → ∃𝑥(𝜑 ∧ 𝜓 ∧ 𝜒)) | ||
Theorem | moexex 2679 | "At most one" double quantification. (Contributed by NM, 3-Dec-2001.) (Proof shortened by Wolf Lammen, 28-Dec-2018.) |
⊢ Ⅎ𝑦𝜑 ⇒ ⊢ ((∃*𝑥𝜑 ∧ ∀𝑥∃*𝑦𝜓) → ∃*𝑦∃𝑥(𝜑 ∧ 𝜓)) | ||
Theorem | moexexv 2680* | "At most one" double quantification. (Contributed by NM, 26-Jan-1997.) |
⊢ ((∃*𝑥𝜑 ∧ ∀𝑥∃*𝑦𝜓) → ∃*𝑦∃𝑥(𝜑 ∧ 𝜓)) | ||
Theorem | 2moex 2681 | Double quantification with "at most one." (Contributed by NM, 3-Dec-2001.) |
⊢ (∃*𝑥∃𝑦𝜑 → ∀𝑦∃*𝑥𝜑) | ||
Theorem | 2euex 2682 | Double quantification with existential uniqueness. (Contributed by NM, 3-Dec-2001.) (Proof shortened by Andrew Salmon, 9-Jul-2011.) |
⊢ (∃!𝑥∃𝑦𝜑 → ∃𝑦∃!𝑥𝜑) | ||
Theorem | 2eumo 2683 | Double quantification with existential uniqueness and "at most one." (Contributed by NM, 3-Dec-2001.) |
⊢ (∃!𝑥∃*𝑦𝜑 → ∃*𝑥∃!𝑦𝜑) | ||
Theorem | 2eu2ex 2684 | Double existential uniqueness. (Contributed by NM, 3-Dec-2001.) |
⊢ (∃!𝑥∃!𝑦𝜑 → ∃𝑥∃𝑦𝜑) | ||
Theorem | 2moswap 2685 | A condition allowing swap of "at most one" and existential quantifiers. (Contributed by NM, 10-Apr-2004.) |
⊢ (∀𝑥∃*𝑦𝜑 → (∃*𝑥∃𝑦𝜑 → ∃*𝑦∃𝑥𝜑)) | ||
Theorem | 2euswap 2686 | A condition allowing swap of uniqueness and existential quantifiers. (Contributed by NM, 10-Apr-2004.) |
⊢ (∀𝑥∃*𝑦𝜑 → (∃!𝑥∃𝑦𝜑 → ∃!𝑦∃𝑥𝜑)) | ||
Theorem | 2exeu 2687 | Double existential uniqueness implies double uniqueness quantification. (Contributed by NM, 3-Dec-2001.) (Proof shortened by Mario Carneiro, 22-Dec-2016.) |
⊢ ((∃!𝑥∃𝑦𝜑 ∧ ∃!𝑦∃𝑥𝜑) → ∃!𝑥∃!𝑦𝜑) | ||
Theorem | 2mo2 2688* | This theorem extends the idea of "at most one" to expressions in two set variables ("at most one pair 𝑥 and 𝑦". Note: this is not expressed by ∃*𝑥∃*𝑦𝜑). 2eu4 2694 relates this extension to double existential uniqueness, if at least one pair exists. (Contributed by Wolf Lammen, 26-Oct-2019.) |
⊢ ((∃*𝑥∃𝑦𝜑 ∧ ∃*𝑦∃𝑥𝜑) ↔ ∃𝑧∃𝑤∀𝑥∀𝑦(𝜑 → (𝑥 = 𝑧 ∧ 𝑦 = 𝑤))) | ||
Theorem | 2mo 2689* | Two equivalent expressions for double "at most one." (Contributed by NM, 2-Feb-2005.) (Revised by Mario Carneiro, 17-Oct-2016.) (Proof shortened by Wolf Lammen, 2-Nov-2019.) |
⊢ (∃𝑧∃𝑤∀𝑥∀𝑦(𝜑 → (𝑥 = 𝑧 ∧ 𝑦 = 𝑤)) ↔ ∀𝑥∀𝑦∀𝑧∀𝑤((𝜑 ∧ [𝑧 / 𝑥][𝑤 / 𝑦]𝜑) → (𝑥 = 𝑧 ∧ 𝑦 = 𝑤))) | ||
Theorem | 2mos 2690* | Double "exists at most one", using implicit substitution. (Contributed by NM, 10-Feb-2005.) |
⊢ ((𝑥 = 𝑧 ∧ 𝑦 = 𝑤) → (𝜑 ↔ 𝜓)) ⇒ ⊢ (∃𝑧∃𝑤∀𝑥∀𝑦(𝜑 → (𝑥 = 𝑧 ∧ 𝑦 = 𝑤)) ↔ ∀𝑥∀𝑦∀𝑧∀𝑤((𝜑 ∧ 𝜓) → (𝑥 = 𝑧 ∧ 𝑦 = 𝑤))) | ||
Theorem | 2eu1 2691 | Double existential uniqueness. This theorem shows a condition under which a "naive" definition matches the correct one. (Contributed by NM, 3-Dec-2001.) (Proof shortened by Wolf Lammen, 11-Nov-2019.) |
⊢ (∀𝑥∃*𝑦𝜑 → (∃!𝑥∃!𝑦𝜑 ↔ (∃!𝑥∃𝑦𝜑 ∧ ∃!𝑦∃𝑥𝜑))) | ||
Theorem | 2eu2 2692 | Double existential uniqueness. (Contributed by NM, 3-Dec-2001.) |
⊢ (∃!𝑦∃𝑥𝜑 → (∃!𝑥∃!𝑦𝜑 ↔ ∃!𝑥∃𝑦𝜑)) | ||
Theorem | 2eu3 2693 | Double existential uniqueness. (Contributed by NM, 3-Dec-2001.) |
⊢ (∀𝑥∀𝑦(∃*𝑥𝜑 ∨ ∃*𝑦𝜑) → ((∃!𝑥∃!𝑦𝜑 ∧ ∃!𝑦∃!𝑥𝜑) ↔ (∃!𝑥∃𝑦𝜑 ∧ ∃!𝑦∃𝑥𝜑))) | ||
Theorem | 2eu4 2694* | This theorem provides us with a definition of double existential uniqueness ("exactly one 𝑥 and exactly one 𝑦"). Naively one might think (incorrectly) that it could be defined by ∃!𝑥∃!𝑦𝜑. See 2eu1 2691 for a condition under which the naive definition holds and 2exeu 2687 for a one-way implication. See 2eu5 2695 and 2eu8 2698 for alternate definitions. (Contributed by NM, 3-Dec-2001.) (Proof shortened by Wolf Lammen, 14-Sep-2019.) |
⊢ ((∃!𝑥∃𝑦𝜑 ∧ ∃!𝑦∃𝑥𝜑) ↔ (∃𝑥∃𝑦𝜑 ∧ ∃𝑧∃𝑤∀𝑥∀𝑦(𝜑 → (𝑥 = 𝑧 ∧ 𝑦 = 𝑤)))) | ||
Theorem | 2eu5 2695* | An alternate definition of double existential uniqueness (see 2eu4 2694). A mistake sometimes made in the literature is to use ∃!𝑥∃!𝑦 to mean "exactly one 𝑥 and exactly one 𝑦." (For example, see Proposition 7.53 of [TakeutiZaring] p. 53.) It turns out that this is actually a weaker assertion, as can be seen by expanding out the formal definitions. This theorem shows that the erroneous definition can be repaired by conjoining ∀𝑥∃*𝑦𝜑 as an additional condition. The correct definition apparently has never been published. (∃* means "exists at most one."). (Contributed by NM, 26-Oct-2003.) |
⊢ ((∃!𝑥∃!𝑦𝜑 ∧ ∀𝑥∃*𝑦𝜑) ↔ (∃𝑥∃𝑦𝜑 ∧ ∃𝑧∃𝑤∀𝑥∀𝑦(𝜑 → (𝑥 = 𝑧 ∧ 𝑦 = 𝑤)))) | ||
Theorem | 2eu6 2696* | Two equivalent expressions for double existential uniqueness. (Contributed by NM, 2-Feb-2005.) (Revised by Mario Carneiro, 17-Oct-2016.) (Proof shortened by Wolf Lammen, 2-Oct-2019.) |
⊢ ((∃!𝑥∃𝑦𝜑 ∧ ∃!𝑦∃𝑥𝜑) ↔ ∃𝑧∃𝑤∀𝑥∀𝑦(𝜑 ↔ (𝑥 = 𝑧 ∧ 𝑦 = 𝑤))) | ||
Theorem | 2eu7 2697 | Two equivalent expressions for double existential uniqueness. (Contributed by NM, 19-Feb-2005.) |
⊢ ((∃!𝑥∃𝑦𝜑 ∧ ∃!𝑦∃𝑥𝜑) ↔ ∃!𝑥∃!𝑦(∃𝑥𝜑 ∧ ∃𝑦𝜑)) | ||
Theorem | 2eu8 2698 | Two equivalent expressions for double existential uniqueness. Curiously, we can put ∃! on either of the internal conjuncts but not both. We can also commute ∃!𝑥∃!𝑦 using 2eu7 2697. (Contributed by NM, 20-Feb-2005.) |
⊢ (∃!𝑥∃!𝑦(∃𝑥𝜑 ∧ ∃𝑦𝜑) ↔ ∃!𝑥∃!𝑦(∃!𝑥𝜑 ∧ ∃𝑦𝜑)) | ||
Theorem | exists1 2699* | Two ways to express "only one thing exists." The left-hand side requires only one variable to express this. Both sides are false in set theory; see theorem dtru 5006. (Contributed by NM, 5-Apr-2004.) |
⊢ (∃!𝑥 𝑥 = 𝑥 ↔ ∀𝑥 𝑥 = 𝑦) | ||
Theorem | exists2 2700 | A condition implying that at least two things exist. (Contributed by NM, 10-Apr-2004.) (Proof shortened by Andrew Salmon, 9-Jul-2011.) |
⊢ ((∃𝑥𝜑 ∧ ∃𝑥 ¬ 𝜑) → ¬ ∃!𝑥 𝑥 = 𝑥) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |