Enderton. Exercises 3.19-3.23.

finite-set-exercises
Joshua Potter 2023-07-06 15:31:09 -06:00
parent 874bbe203c
commit 1c6fec389f
3 changed files with 257 additions and 44 deletions

View File

@ -75,7 +75,7 @@ The \textbf{composition} of sets $F$ and $G$ is
\begin{definition} \begin{definition}
\lean*{Common/Set/Relation}{Set.Relation.comp} \lean*{Bookshelf/Enderton/Set/Relation}{Set.Relation.comp}
\end{definition} \end{definition}
@ -87,7 +87,7 @@ The \textbf{domain} of set $R$, denoted $\dom{R}$, is given by
\begin{definition} \begin{definition}
\lean*{Common/Set/Relation}{Set.Relation.dom} \lean*{Bookshelf/Enderton/Set/Relation}{Set.Relation.dom}
\end{definition} \end{definition}
@ -137,7 +137,7 @@ Given \nameref{ref:relation} $R$, the \textbf{field} of $R$, denoted $\fld{R}$,
\begin{definition} \begin{definition}
\lean*{Common/Set/Relation}{Set.Relation.fld} \lean*{Bookshelf/Enderton/Set/Relation}{Set.Relation.fld}
\end{definition} \end{definition}
@ -182,7 +182,7 @@ The \textbf{image of $A$ under $F$} is the set
\begin{definition} \begin{definition}
\lean*{Common/Set/Relation}{Set.Relation.image} \lean*{Bookshelf/Enderton/Set/Relation}{Set.Relation.image}
\end{definition} \end{definition}
@ -194,7 +194,7 @@ The \textbf{inverse} of a set $F$ is the set
\begin{definition} \begin{definition}
\lean*{Common/Set/Relation}{Set.Relation.inv} \lean*{Bookshelf/Enderton/Set/Relation}{Set.Relation.inv}
\end{definition} \end{definition}
@ -206,7 +206,7 @@ For any sets $u$ and $v$, the \textbf{ordered pair} $\left< u, v \right>$ is
\begin{definition} \begin{definition}
\lean*{Common/Set/OrderedPair}{OrderedPair} \lean*{Bookshelf/Enderton/Set/OrderedPair}{OrderedPair}
\end{definition} \end{definition}
@ -305,7 +305,7 @@ The \textbf{range} of set $R$, denoted $\ran{R}$, is given by
\begin{definition} \begin{definition}
\lean*{Common/Set/Relation}{Set.Relation.ran} \lean*{Bookshelf/Enderton/Set/Relation}{Set.Relation.ran}
\end{definition} \end{definition}
@ -328,7 +328,7 @@ A \textbf{relation} is a set of \nameref{ref:ordered-pair}s.
\begin{definition} \begin{definition}
\lean*{Common/Set/Relation}{Set.Relation} \lean*{Bookshelf/Enderton/Set/Relation}{Set.Relation}
\end{definition} \end{definition}
@ -340,7 +340,7 @@ The \textbf{restriction} of a set $F$ to set $A$ is the set
\begin{definition} \begin{definition}
\lean*{Common/Set/Relation}{Set.Relation.restriction} \lean*{Bookshelf/Enderton/Set/Relation}{Set.Relation.restriction}
\end{definition} \end{definition}
@ -1341,7 +1341,8 @@ Let $A$ and $B$ be sets. $x \not\in A + B$ if and only if either
\begin{proof} \begin{proof}
\lean{Common/Set/Basic}{Set.not\_mem\_symm\_diff\_inter\_or\_not\_union} \lean{Bookshelf/Enderton/Set/Basic}
{Set.not\_mem\_symm\_diff\_inter\_or\_not\_union}
By definition of the \nameref{ref:symmetric-difference}, By definition of the \nameref{ref:symmetric-difference},
\begin{align*} \begin{align*}
@ -2657,13 +2658,13 @@ If not, then under what conditions does equality hold?
\statementpadding \statementpadding
\lean*{Common/Set/Relation} \lean*{Bookshelf/Enderton/Set/Relation}
{Set.Relation.dom\_inv\_eq\_ran\_self} {Set.Relation.dom\_inv\_eq\_ran\_self}
\lean*{Common/Set/Relation} \lean*{Bookshelf/Enderton/Set/Relation}
{Set.Relation.ran\_inv\_eq\_dom\_self} {Set.Relation.ran\_inv\_eq\_dom\_self}
\lean{Common/Set/Relation} \lean{Bookshelf/Enderton/Set/Relation}
{Set.Relation.inv\_inv\_eq\_self} {Set.Relation.inv\_inv\_eq\_self}
We prove that (i) $\dom{(F^{-1})} = \ran{F}$, (ii) $\ran{(F^{-1})} = \dom{F}$, We prove that (i) $\dom{(F^{-1})} = \ran{F}$, (ii) $\ran{(F^{-1})} = \dom{F}$,
@ -2715,10 +2716,10 @@ If not, then under what conditions does equality hold?
\statementpadding \statementpadding
\lean*{Common/Set/Relation} \lean*{Bookshelf/Enderton/Set/Relation}
{Set.Relation.single\_valued\_inv\_iff\_single\_rooted\_self} {Set.Relation.single\_valued\_inv\_iff\_single\_rooted\_self}
\lean{Common/Set/Relation} \lean{Bookshelf/Enderton/Set/Relation}
{Set.Relation.single\_valued\_self\_iff\_single\_rooted\_inv} {Set.Relation.single\_valued\_self\_iff\_single\_rooted\_inv}
We prove that (i) any set $F$, $F^{-1}$ is a function iff $F$ is We prove that (i) any set $F$, $F^{-1}$ is a function iff $F$ is
@ -2778,7 +2779,7 @@ For any one-to-one function $F$, $F^{-1}$ is also one-to-one.
\begin{proof} \begin{proof}
\lean{Common/Set/Relation} \lean{Bookshelf/Enderton/Set/Relation}
{Set.Relation.one\_to\_one\_self\_iff\_one\_to\_one\_inv} {Set.Relation.one\_to\_one\_self\_iff\_one\_to\_one\_inv}
We prove that (i) $F^{-1}$ is a function and (ii) $F^{-1}$ is single-rooted. We prove that (i) $F^{-1}$ is a function and (ii) $F^{-1}$ is single-rooted.
@ -2868,7 +2869,7 @@ For any one-to-one function $F$, $F^{-1}$ is also one-to-one.
\statementpadding \statementpadding
\lean*{Common/Set/Relation} \lean*{Bookshelf/Enderton/Set/Relation}
{Set.Relation.single\_valued\_comp\_is\_single\_valued} {Set.Relation.single\_valued\_comp\_is\_single\_valued}
\lean{Bookshelf/Enderton/Set/Chapter\_3} \lean{Bookshelf/Enderton/Set/Chapter\_3}
@ -2924,7 +2925,7 @@ For any one-to-one function $F$, $F^{-1}$ is also one-to-one.
\begin{proof} \begin{proof}
\lean{Common/Set/Relation} \lean{Bookshelf/Enderton/Set/Relation}
{Set.Relation.comp\_inv\_eq\_inv\_comp\_inv} {Set.Relation.comp\_inv\_eq\_inv\_comp\_inv}
By definition of the \nameref{ref:composition} of $F$ and $G$, By definition of the \nameref{ref:composition} of $F$ and $G$,
@ -4296,13 +4297,16 @@ Evaluate each of the following: $A(\emptyset)$, $\img{A}{\emptyset}$,
\end{proof} \end{proof}
\subsection{\pending{Exercise 3.20}}% \subsection{\verified{Exercise 3.20}}%
\label{sub:exercise-3.20} \label{sub:exercise-3.20}
Show that $F \restriction A = F \cap (A \times \ran{F})$. Show that $F \restriction A = F \cap (A \times \ran{F})$.
\begin{proof} \begin{proof}
\lean{Bookshelf/Enderton/Set/Chapter\_3}
{Enderton.Set.Chapter\_3.exercise\_3\_20}
Let $F$ and $A$ be arbitrary sets. Let $F$ and $A$ be arbitrary sets.
By \nameref{sub:corollary-3c} and definition of the \nameref{ref:restriction}, By \nameref{sub:corollary-3c} and definition of the \nameref{ref:restriction},
intersection, and \nameref{ref:range} of sets, intersection, and \nameref{ref:range} of sets,
@ -4319,34 +4323,38 @@ Show that $F \restriction A = F \cap (A \times \ran{F})$.
\end{proof} \end{proof}
\subsection{\pending{Exercise 3.21}}% \subsection{\verified{Exercise 3.21}}%
\label{sub:exercise-3.21} \label{sub:exercise-3.21}
Show that $(R \circ S) \circ T = R \circ (S \circ T)$. Show that $(R \circ S) \circ T = R \circ (S \circ T)$.
\begin{proof} \begin{proof}
\lean{Bookshelf/Enderton/Set/Relation}
{Set.Relation.comp\_assoc}
Let $R$, $S$, and $T$ be arbitrary sets. Let $R$, $S$, and $T$ be arbitrary sets.
By definition of the \nameref{ref:composition} of sets, By definition of the \nameref{ref:composition} of sets,
\begin{align*} \begin{align*}
(R \circ S) \circ T (R \circ S) \circ T
& = \{\left< u, v \right> \mid & = \{\left< u, v \right> \mid
\exists t(u(R \circ S)t \land tTv)\} \\ \exists t(uTt \land t(R \circ S)v)\} \\
& = \{\left< u, v \right> \mid & = \{\left< u, v \right> \mid
\exists t((\exists a, uRa \land aSt) \land tTv)\} \\ \exists t(uTt \land (\exists a(tSa \land aRv))\} \\
& = \{\left< u, v \right> \mid & = \{\left< u, v \right> \mid
\exists t, \exists a, (uRa \land aSt \land tTv)\} \\ \exists t, \exists a, (uTt \land tSa) \land aRv)\} \\
& = \{\left< u, v \right> \mid & = \{\left< u, v \right> \mid
\exists a, \exists t, (uRa \land aSt \land tTv)\} \\ \exists a, \exists t, (uTt \land tSa) \land aRv)\} \\
& = \{\left< u, v \right> \mid & = \{\left< u, v \right> \mid
\exists a(uRa \land (\exists t, aSt \land tTv))\} \\ \exists a, (\exists t(uTt \land tSa)) \land aRv)\} \\
& = \{\left< u, v \right> \mid \exists a(uRa \land a(S \circ T)v)\} \\ & = \{\left< u, v \right> \mid
\exists a, u(S \circ T)a \land aRv)\} \\
& = R \circ (S \circ T). & = R \circ (S \circ T).
\end{align*} \end{align*}
\end{proof} \end{proof}
\subsection{\pending{Exercise 3.22}}% \subsection{\verified{Exercise 3.22}}%
\label{sub:exercise-3.22} \label{sub:exercise-3.22}
Show that the following are correct for any sets. Show that the following are correct for any sets.
@ -4360,6 +4368,17 @@ Show that the following are correct for any sets.
\begin{proof} \begin{proof}
\statementpadding
\lean*{Bookshelf/Enderton/Set/Relation}
{Enderton.Set.Chapter\_3.exercise\_3\_22\_a}
\lean*{Bookshelf/Enderton/Set/Relation}
{Enderton.Set.Chapter\_3.exercise\_3\_22\_b}
\lean{Bookshelf/Enderton/Set/Relation}
{Enderton.Set.Chapter\_3.exercise\_3\_22\_c}
Let $A$, $B$, $F$, $G$, and $Q$ be arbitrary sets. Let $A$, $B$, $F$, $G$, and $Q$ be arbitrary sets.
\paragraph{(a)}% \paragraph{(a)}%
@ -4408,7 +4427,7 @@ Show that the following are correct for any sets.
\end{proof} \end{proof}
\subsection{\pending{Exercise 3.23}}% \subsection{\verified{Exercise 3.23}}%
\label{sub:exercise-3.23} \label{sub:exercise-3.23}
Let $I_A$ be the identity function on the set $A$. Let $I_A$ be the identity function on the set $A$.
@ -4418,6 +4437,14 @@ Show that for any sets $B$ and $C$,
\begin{proof} \begin{proof}
\statementpadding
\lean*{Bookshelf/Enderton/Set/Relation}
{Enderton.Set.Chapter\_3.exercise\_3\_23\_i}
\lean{Bookshelf/Enderton/Set/Relation}
{Enderton.Set.Chapter\_3.exercise\_3\_23\_ii}
Let $I_A$ be the identity function on the set $A$. Let $I_A$ be the identity function on the set $A$.
That is, $I_A = \{\left< u, u \right> \mid u \in A\}$. That is, $I_A = \{\left< u, u \right> \mid u \in A\}$.
Let $B$ and $C$ be any sets. Let $B$ and $C$ be any sets.
@ -4433,8 +4460,8 @@ Show that for any sets $B$ and $C$,
Let $\left< x, y \right> \in B \circ I_A$. Let $\left< x, y \right> \in B \circ I_A$.
By definition of the \nameref{ref:composition} of sets, By definition of the \nameref{ref:composition} of sets,
there exists some $t$ such that $xBt$ and $t(I_A)y$. there exists some $t$ such that $x(I_A)t$ and $tBy$.
By definition of the identity function, $I_A(t) = y$ implies $t = y$. By definition of the identity function, $I_A(x) = t$ implies $x = t$.
Thus $xBy$. Thus $xBy$.
By hypothesis, $x \in \dom{(B \circ I_A)}$. By hypothesis, $x \in \dom{(B \circ I_A)}$.
Therefore $x \in \dom{I_A} = A$. Therefore $x \in \dom{I_A} = A$.

View File

@ -1296,6 +1296,153 @@ theorem exercise_3_19_x
end Exercise_3_19 end Exercise_3_19
/-- #### Exercise 3.20
Show that `F ↾ A = F ∩ (A × ran F)`.
-/
theorem exercise_3_20 {F : Set.Relation α} {A : Set α}
: F.restriction A = F ∩ (Set.prod A (ran F)) := by
calc F.restriction A
_ = {p | p ∈ F ∧ p.fst ∈ A} := rfl
_ = {p | p ∈ F ∧ p.fst ∈ A ∧ p.snd ∈ ran F} := by
ext x
have (a, b) := x
simp only [
Set.mem_setOf_eq, Set.sep_and, Set.mem_inter_iff, iff_self_and, and_imp
]
intro hF _
exact ⟨hF, mem_pair_imp_snd_mem_ran hF⟩
_ = {p | p ∈ F} ∩ {p | p.fst ∈ A ∧ p.snd ∈ ran F} := rfl
_ = F ∩ {p | p.fst ∈ A ∧ p.snd ∈ ran F} := rfl
_ = F ∩ (Set.prod A (ran F)) := rfl
/-- #### Exercise 3.22 (a)
Show that the following is correct for any sets.
```
A ⊆ B → F⟦A⟧ ⊆ F⟦B⟧
```
-/
theorem exercise_3_22_a {A B : Set α} {F : Set.Relation α} (h : A ⊆ B)
: F.image A ⊆ F.image B := by
show ∀ x, x ∈ F.image A → x ∈ F.image B
unfold image
simp only [Set.mem_setOf_eq]
intro x hx
have ⟨u, hu⟩ := hx
have := h hu.left
exact ⟨u, this, hu.right⟩
/-- #### Exercise 3.22 (b)
Show that the following is correct for any sets.
```
(F ∘ G)⟦A⟧ = F⟦G⟦A⟧⟧
```
-/
theorem exercise_3_22_b {A B : Set α} {F : Set.Relation α}
: (F.comp G).image A = F.image (G.image A) := by
calc (F.comp G).image A
_ = { v | ∃ u ∈ A, (u, v) ∈ F.comp G } := rfl
_ = { v | ∃ u ∈ A, ∃ a, (u, a) ∈ G ∧ (a, v) ∈ F } := rfl
_ = { v | ∃ a, ∃ u ∈ A, (u, a) ∈ G ∧ (a, v) ∈ F } := by
ext p
simp only [Set.mem_setOf_eq]
apply Iff.intro
· intro ⟨u, hu, a, ha⟩
exact ⟨a, u, hu, ha⟩
· intro ⟨a, u, hu, ha⟩
exact ⟨u, hu, a, ha⟩
_ = { v | ∃ a, (∃ u ∈ A, (u, a) ∈ G) ∧ (a, v) ∈ F } := by
ext p
simp only [Set.mem_setOf_eq]
apply Iff.intro
· intro ⟨a, u, h⟩
exact ⟨a, ⟨u, h.left, h.right.left⟩, h.right.right⟩
· intro ⟨a, ⟨u, hu⟩, ha⟩
exact ⟨a, u, hu.left, hu.right, ha⟩
_ = { v | ∃ a ∈ { w | ∃ u ∈ A, (u, w) ∈ G }, (a, v) ∈ F } := rfl
_ = { v | ∃ a ∈ G.image A, (a, v) ∈ F } := rfl
_ = F.image (G.image A) := rfl
/-- #### Exercise 3.22 (c)
Show that the following is correct for any sets.
```
Q ↾ (A B) = (Q ↾ A) (Q ↾ B)
```
-/
theorem exercise_3_22_c {A B : Set α} {Q : Set.Relation α}
: Q.restriction (A B) = (Q.restriction A) (Q.restriction B) := by
calc Q.restriction (A B)
_ = { p | p ∈ Q ∧ p.1 ∈ A B } := rfl
_ = { p | p ∈ Q ∧ (p.1 ∈ A p.1 ∈ B) } := rfl
_ = { p | (p ∈ Q ∧ p.1 ∈ A) (p ∈ Q ∧ p.1 ∈ B) } := by
ext p
simp only [Set.sep_or, Set.mem_union, Set.mem_setOf_eq]
_ = { p | p ∈ Q ∧ p.1 ∈ A} { p | p ∈ Q ∧ p.1 ∈ B } := rfl
_ = (Q.restriction A) (Q.restriction B) := rfl
/-- #### Exercise 3.23 (i)
Let `I` be the identity function on the set `A`. Show that for any sets `B` and
`C`, `B ∘ I = B ↾ A`.
-/
theorem exercise_3_23_i {A : Set α} {B : Set.Relation α} {I : Set.Relation α}
(hI : I = { p | p.1 ∈ A ∧ p.1 = p.2 })
: B.comp I = B.restriction A := by
rw [Set.Subset.antisymm_iff]
apply And.intro
· show ∀ p, p ∈ B.comp I → p ∈ B.restriction A
intro (x, y) hp
have ⟨t, ht⟩ := hp
rw [hI] at ht
simp only [Set.mem_setOf_eq] at ht
show (x, y) ∈ B ∧ x ∈ A
rw [← ht.left.right] at ht
exact ⟨ht.right, ht.left.left⟩
· show ∀ p, p ∈ B.restriction A → p ∈ B.comp I
unfold restriction comp
rw [hI]
simp only [Set.mem_setOf_eq, and_true]
intro (x, y) hp
refine ⟨x, ⟨hp.right, rfl⟩, hp.left⟩
/-- #### Exercise 3.23 (ii)
Let `I` be the identity function on the set `A`. Show that for any sets `B` and
`C`, `I⟦C⟧ = A ∩ C`.
-/
theorem exercise_3_23_ii {A C : Set α} {I : Set.Relation α}
(hI : I = { p | p.1 ∈ A ∧ p.1 = p.2 })
: I.image C = A ∩ C := by
calc I.image C
_ = { v | ∃ u ∈ C, (u, v) ∈ I } := rfl
_ = { v | ∃ u ∈ C, u ∈ A ∧ u = v } := by
ext v
simp only [Set.mem_setOf_eq]
apply Iff.intro
· intro ⟨u, h₁, h₂⟩
rw [hI] at h₂
simp only [Set.mem_setOf_eq] at h₂
exact ⟨u, h₁, h₂⟩
· intro ⟨u, h₁, h₂⟩
refine ⟨u, h₁, ?_⟩
· rw [hI]
simp only [Set.mem_setOf_eq]
exact h₂
_ = { v | v ∈ C ∧ v ∈ A } := by
ext v
simp only [Set.mem_setOf_eq, Set.sep_mem_eq, Set.mem_inter_iff]
apply Iff.intro
· intro ⟨u, hC, hA, hv⟩
rw [← hv]
exact ⟨hC, hA⟩
· intro ⟨hC, hA⟩
exact ⟨v, hC, hA, rfl⟩
_ = C ∩ A := rfl
_ = A ∩ C := Set.inter_comm C A
end Relation end Relation
end Enderton.Set.Chapter_3 end Enderton.Set.Chapter_3

View File

@ -278,6 +278,8 @@ theorem single_valued_subset {F G : Set.Relation α}
intro y₁ hy₁ intro y₁ hy₁
exact single_valued_eq_unique hG (h hy₁.right) (h hy) exact single_valued_eq_unique hG (h hy₁.right) (h hy)
/-! ## Injections -/
/-- /--
A `Relation` `R` is one-to-one if it is a single-rooted function. A `Relation` `R` is one-to-one if it is a single-rooted function.
-/ -/
@ -300,6 +302,22 @@ theorem one_to_one_self_iff_one_to_one_inv {R : Relation α}
· intro ⟨hx, hy⟩ · intro ⟨hx, hy⟩
exact ⟨hy, hx⟩ exact ⟨hy, hx⟩
/-! ## Surjections -/
/--
Indicates `Relation` `F` is a function from `A` to `B`.
This is usually denoted as `F : A → B`.
-/
def mapsInto (F : Relation α) (A B : Set α) :=
F.isSingleValued ∧ dom F = A ∧ ran F ⊆ B
/--
Indicates `Relation` `F` is a function from `A` to `ran F = B`.
-/
def mapsOnto (F : Relation α) (A B : Set α) :=
F.isSingleValued ∧ dom F = A ∧ ran F = B
/-! ## Composition -/ /-! ## Composition -/
/-- /--
@ -341,6 +359,41 @@ theorem ran_comp_imp_ran_self {F G : Relation α}
intro x t ht intro x t ht
exact ⟨t, ht.right⟩ exact ⟨t, ht.right⟩
/--
Composition of functions is associative.
-/
theorem comp_assoc {R S T : Relation α}
: (R.comp S).comp T = R.comp (S.comp T) := by
calc (R.comp S).comp T
_ = { p | ∃ t, (p.1, t) ∈ T ∧ (t, p.2) ∈ R.comp S} := rfl
_ = { p | ∃ t, (p.1, t) ∈ T ∧ (∃ a, (t, a) ∈ S ∧ (a, p.2) ∈ R) } := rfl
_ = { p | ∃ t, ∃ a, ((p.1, t) ∈ T ∧ (t, a) ∈ S) ∧ (a, p.2) ∈ R } := by
ext p
simp only [mem_setOf_eq]
apply Iff.intro
· intro ⟨t, ht, a, ha⟩
exact ⟨t, a, ⟨ht, ha.left⟩, ha.right⟩
· intro ⟨t, a, h₁, h₂⟩
exact ⟨t, h₁.left, a, h₁.right, h₂⟩
_ = { p | ∃ a, ∃ t, ((p.1, t) ∈ T ∧ (t, a) ∈ S) ∧ (a, p.2) ∈ R } := by
ext p
simp only [mem_setOf_eq]
apply Iff.intro
· intro ⟨t, a, h⟩
exact ⟨a, t, h⟩
· intro ⟨a, t, h⟩
exact ⟨t, a, h⟩
_ = { p | ∃ a, (∃ t, (p.1, t) ∈ T ∧ (t, a) ∈ S) ∧ (a, p.2) ∈ R } := by
ext p
simp only [mem_setOf_eq]
apply Iff.intro
· intro ⟨a, t, h⟩
exact ⟨a, ⟨t, h.left⟩, h.right⟩
· intro ⟨a, ⟨t, ht⟩, ha⟩
exact ⟨a, t, ht, ha⟩
_ = { p | ∃ a, (p.1, a) ∈ S.comp T ∧ (a, p.2) ∈ R } := rfl
_ = R.comp (S.comp T) := rfl
/-- /--
The composition of two functions is itself a function. The composition of two functions is itself a function.
-/ -/
@ -420,20 +473,6 @@ theorem comp_inv_eq_inv_comp_inv {F G : Relation α}
refine ⟨t, ?_, ?_⟩ <;> rwa [← mem_self_comm_mem_inv] refine ⟨t, ?_, ?_⟩ <;> rwa [← mem_self_comm_mem_inv]
_ = comp (inv G) (inv F) := rfl _ = comp (inv G) (inv F) := rfl
/--
Indicates `Relation` `F` is a function from `A` to `B`.
This is usually denoted as `F : A → B`.
-/
def mapsInto (F : Relation α) (A B : Set α) :=
F.isSingleValued ∧ dom F = A ∧ ran F ⊆ B
/--
Indicates `Relation` `F` is a function from `A` to `ran F = B`.
-/
def mapsOnto (F : Relation α) (A B : Set α) :=
F.isSingleValued ∧ dom F = A ∧ ran F = B
/-! ## Ordered Pairs -/ /-! ## Ordered Pairs -/
/-- /--