Description: Define restricted existential uniqueness.
Note: This notation is most often used to express that ph holds for exactly one element of a given class A . For this reading F/_ x A is required, though, for example, asserted when x and A are disjoint.
Should instead A depend on x , you rather assert exactly one x fulfilling ph happens to be contained in the corresponding A ( x ) . This interpretation is rarely needed (see also df-ral ). (Contributed by NM, 22-Nov-1994)
| Ref | Expression | ||
|---|---|---|---|
| Assertion | df-reu | ⊢ ( ∃! 𝑥 ∈ 𝐴 𝜑 ↔ ∃! 𝑥 ( 𝑥 ∈ 𝐴 ∧ 𝜑 ) ) | 
| Step | Hyp | Ref | Expression | 
|---|---|---|---|
| 0 | vx | ⊢ 𝑥 | |
| 1 | cA | ⊢ 𝐴 | |
| 2 | wph | ⊢ 𝜑 | |
| 3 | 2 0 1 | wreu | ⊢ ∃! 𝑥 ∈ 𝐴 𝜑 | 
| 4 | 0 | cv | ⊢ 𝑥 | 
| 5 | 4 1 | wcel | ⊢ 𝑥 ∈ 𝐴 | 
| 6 | 5 2 | wa | ⊢ ( 𝑥 ∈ 𝐴 ∧ 𝜑 ) | 
| 7 | 6 0 | weu | ⊢ ∃! 𝑥 ( 𝑥 ∈ 𝐴 ∧ 𝜑 ) | 
| 8 | 3 7 | wb | ⊢ ( ∃! 𝑥 ∈ 𝐴 𝜑 ↔ ∃! 𝑥 ( 𝑥 ∈ 𝐴 ∧ 𝜑 ) ) |