diff --git a/Bookshelf/Enderton/Set.tex b/Bookshelf/Enderton/Set.tex index 907bf4d..cbe9530 100644 --- a/Bookshelf/Enderton/Set.tex +++ b/Bookshelf/Enderton/Set.tex @@ -112,6 +112,18 @@ For each formula $\phi$ not containing $B$, the following is an axiom: \end{axiom} +\section{\defined{Symmetric Difference}}% +\label{ref:symmetric-difference} + +The \textbf{symmetric difference} $A + B$ of sets $A$ and $B$ is the set + $(A - B) \cup (B - A)$. + +\begin{definition} + + \lean*{Mathlib/Data/Set/Basic}{symmDiff\_def} + +\end{definition} + \section{\defined{Union Axiom}}% \label{ref:union-axiom} @@ -1404,10 +1416,64 @@ For any set $C$ and $\mathscr{A} \neq \emptyset$, \end{proof} +\subsection{\verified{% + \texorpdfstring{$\cap$/$-$}{Intersection/Difference} Associativity}}% +\label{sub:intersection-difference-associativity} + +Let $A$, $B$, and $C$ be sets. +Then $A \cap (B - C) = (A \cap B) - C$. + +\begin{proof} + + \lean{Mathlib/Data/Set/Basic}{Set.inter\_diff\_assoc} + + Let $A$, $B$, and $C$ be sets. + By definition of the intersection and relative complement of sets, + \begin{align*} + A \cap (B - C) + & = \{ x \mid x \in A \land x \in B - C \} \\ + & = \{ x \mid x \in A \land (x \in B \land x \not\in C) \} \\ + & = \{ x \mid (x \in A \land x \in B) \land x \not\in C \} \\ + & = \{ x \mid x \in A \cap B \land x \not \in C \} \\ + & = (A \cap B) - C. + \end{align*} + +\end{proof} + +\subsection{\verified{Nonmembership of Symmetric Difference}} +\label{sub:nonmembership-symmetric-difference} + +Let $A$ and $B$ be sets. $x \not\in A + B$ if and only if either + $x \in A \cap B$ or $x \not\in A \cup B$. + +\begin{proof} + + \lean{Common/Set/Basic}{Set.not\_mem\_symm\_diff\_inter\_or\_not\_union} + + By definition of the \nameref{ref:symmetric-difference}, + \begin{align*} + x \not\in A + B + & = \neg(x \in A + B) \\ + & = \neg[x \in (A - B) \cup (B - A)] \\ + & = \neg[x \in (A - B) \lor x \in (B - A)] \\ + & = \neg[(x \in A \land x \not\in B) \lor + (x \in B \land x \not\in A)] \\ + & = \neg(x \in A \land x \not\in B) \land + \neg(x \in B \land x \not\in A) \\ + & = (x \not\in A \lor x \in B) \land (x \not\in B \lor x \in A) \\ + & = ((x \not\in A \lor x \in B) \land x \not\in B) \lor + ((x \not\in A \lor x \in B) \land x \in A) \\ + & = (x \not\in A \land x \not\in B) \lor (x \in B \land x \in A) \\ + & = \neg(x \in A \lor x \in B) \lor (x \in B \land x \in A) \\ + & = x \not\in A \cup B \text{ or } x \in A \cap B. + \end{align*} + +\end{proof} + \section{Exercises 4}% \label{sec:exercises-4} -\subsection{\unverified{Exercise 4.11}}% +\subsection{\verified{Exercise 4.11}}% \label{sub:exercise-4.11} Show that for any sets $A$ and $B$, @@ -1416,11 +1482,60 @@ Show that for any sets $A$ and $B$, \begin{proof} - TODO + \statementpadding + + \lean*{Bookshelf/Enderton/Set/Chapter\_2} + {Enderton.Set.Chapter\_2.exercise\_4\_11\_i} + + \lean{Bookshelf/Enderton/Set/Chapter\_2} + {Enderton.Set.Chapter\_2.exercise\_4\_11\_ii} + + \noindent Let $A$ and $B$ be sets. + We prove that + \begin{enumerate}[(i)] + \item $A = (A \cap B) \cup (A - B)$ + \item $A \cup (B - A) = A \cup B$ + \end{enumerate} + + \paragraph{(i)}% + + By definition of the intersection, union, and relative complements of sets, + \begin{align*} + (A \cap B) \cup (A - B) + & = \{ x \mid x \in A \cap B \lor x \in A - B \} \\ + & = \{ x \mid x \in \{ y \mid y \in A \land y \in B \} \lor + x \in A - B \} \\ + & = \{ x \mid (x \in A \land x \in B) \lor x \in A - B \} \\ + & = \{ x \mid (x \in A \land x \in B) \lor + x \in \{ y \mid y \in A \land y \not\in B \} \} \\ + & = \{ x \mid (x \in A \land x \in B) \lor + (x \in A \land x \not\in B) \} \\ + & = \{ x \mid x \in A \lor (x \in B \land x \not\in B) \} \\ + & = \{ x \mid x \in A \lor F \} \\ + & = \{ x \mid x \in A \} \\ + & = A. + \end{align*} + + \paragraph{(ii)}% + + By definition of the union and relative complements of sets, + \begin{align*} + A \cup (B - A) + & = \{ x \mid x \in A \lor x \in B - A \} \\ + & = \{ x \mid x \in A \lor + x \in \{ y \mid y \in B \land y \not\in A \} \} \\ + & = \{ x \mid x \in A \lor (x \in B \land x \not\in A) \} \\ + & = \{ x \mid (x \in A \lor x \in B) \land + (x \in A \lor x \not\in A) \} \\ + & = \{ x \mid (x \in A \lor x \in B) \land T \} \\ + & = \{ x \mid x \in A \lor x \in B \} \\ + & = \{ x \mid x \in A \cup B \} \\ + & = A \cup B. + \end{align*} \end{proof} -\subsection{\unverified{Exercise 4.12}}% +\subsection{\verified{Exercise 4.12}}% \label{sub:exercise-4.12} Verify the following identity (one of De Morgan's laws): @@ -1428,22 +1543,22 @@ Verify the following identity (one of De Morgan's laws): \begin{proof} - TODO + Refer to \nameref{sub:de-morgans-laws}. \end{proof} -\subsection{\unverified{Exercise 4.13}}% +\subsection{\verified{Exercise 4.13}}% \label{sub:exercise-4.13} Show that if $A \subseteq B$, then $C - B \subseteq C - A$. \begin{proof} - TODO + Refer to \nameref{sub:anti-monotonicity}. \end{proof} -\subsection{\unverified{Exercise 4.14}}% +\subsection{\verified{Exercise 4.14}}% \label{sub:exercise-4.14} Show by example that for some sets $A$, $B$, and $C$, the set $A - (B - C)$ is @@ -1451,35 +1566,157 @@ Show by example that for some sets $A$, $B$, and $C$, the set $A - (B - C)$ is \begin{proof} - TODO + \lean{Bookshelf/Enderton/Set/Chapter\_2} + {Enderton.Set.Chapter\_2.exercise\_4\_14} + + Let $A = \{1, 2, 3\}$, $B = \{2, 3, 4\}$, and $C = \{3, 4, 5\}$. + Then + \begin{align*} + A - (B - C) + & = \{1, 2, 3\} - (\{2, 3, 4\} - \{3, 4, 5\}) \\ + & = \{1, 2, 3\} - \{2\} \\ + & = \{1, 3\} + \end{align*} + but + \begin{align*} + (A - B) - C + & = (\{1, 2, 3\} - \{2, 3, 4\}) - \{3, 4, 5\} \\ + & = \{1\} - \{3, 4, 5\} \\ + & = \{1\}. + \end{align*} \end{proof} -\subsection{\unverified{Exercise 4.15}}% -\label{sub:exercise-4.15} - -Define the symmetric difference $A + B$ of sets $A$ and $B$ to be the set - $(A - B) \cup (B - A)$. - -\subsubsection{\unverified{Exercise 4.15a}}% -\label{ssub:exercise-4.15a} +\subsection{\verified{Exercise 4.15a}}% +\label{sub:exercise-4.15a} Show that $A \cap (B + C) = (A \cap B) + (A \cap C)$. \begin{proof} - TODO + \lean{Mathlib/Data/Set/Basic}{Set.inter\_symmDiff\_distrib\_left} + + By definition of the intersection, \nameref{ref:symmetric-difference}, and + relative complement of sets, + \begin{align*} + (A & \cap B) + (A \cap C) \\ + & = [(A \cap B) - (A \cap C)] \cup [(A \cap C) - (A \cap B)] \\ + & = [(A \cap B) - A] \\ + & \qquad \cup [(A \cap B) - C] \\ + & \qquad \cup [(A \cap C) - A] \\ + & \qquad \cup [(A \cap C) - B] + & \textref{sub:de-morgans-laws} \\ + & = [A \cap (B - A)] \\ + & \qquad \cup [A \cap (B - C)] \\ + & \qquad \cup [A \cap (C - A)] \\ + & \qquad \cup [A \cap (C - B)] + & \textref{sub:intersection-difference-associativity} \\ + & = \emptyset \\ + & \qquad \cup [A \cap (B - C)] \\ + & \qquad \cup \emptyset \\ + & \qquad \cup [A \cap (C - B)] + & \textref{sub:identitives-involving-empty-set} \\ + & = [A \cap (B - C)] \cup [A \cap (C - B)] \\ + & = A \cap [(B - C) \cup (C - B)] + & \textref{sub:distributive-laws} \\ + & = A \cap (B + C). + \end{align*} \end{proof} -\subsubsection{\unverified{Exercise 4.15b}}% -\label{ssub:exercise-4.15b} +\subsection{\verified{Exercise 4.15b}}% +\label{sub:exercise-4.15b} Show that $A + (B + C) = (A + B) + C$. \begin{proof} - TODO + \lean{Mathlib/Data/Set/Basic}{Set.symmDiff\_assoc} + + \noindent Let $A$, $B$, and $C$ be sets. + We prove that + \begin{enumerate}[(i)] + \item $A + (B + C) \subseteq (A + B) + C$ + \item $(A + B) + C \subseteq A + (B + C)$ + \end{enumerate} + + \paragraph{(i)}% + \label{par:exercise-4.15b-i} + + Let $x \in A + (B + C)$. + Then $x$ is in $A$ or in $B + C$, but not both. + There are two cases to consider: + + \subparagraph{Case 1}% + + Suppose $x \in A$ and $x \not\in B + C$. + Then, by \nameref{sub:nonmembership-symmetric-difference}, + (a) $x \in B \cap C$ or (b) $x \not\in B \cup C$. + Suppose (a) was true. + That is, $x \in B$ and $x \in C$. + Since $x$ is a member of $A$ and $B$, $x \not\in (A + B)$. + Since $x$ is not a member of $A + B$ but is a member of $C$, + $x \in (A + B) + C$. + Now suppose (b) was true. + That is, $x \not\in B$ and $x \not\in C$. + Since $x$ is a member of $A$ but not $B$, $x \in (A + B)$. + Since $x$ is a member of $A + B$ but not $C$, $x \in (A + B) + C$. + + \subparagraph{Case 2}% + + Suppose $x \in B + C$ and $x \not\in A$. + Then (a) $x \in B$ or (b) $x \in C$ but not both. + Suppose (a) was true. + That is, $x \in B$ and $x \not\in C$. + Since $x$ is not a member of $A$ and is a member of $B$, $x \in A + B$. + Since $x$ is a member of $A + B$ but not $C$, $x \in (A + B) + C$. + Now suppose (b) was true. + That is, $x \not\in B$ and $x \in C$. + Since $x$ is not a member of $A$ nor a member of $B$, $x \not\in A + B$. + Since $x$ is not a member of $A + B$ but is a member of $C$, + $x \in (A + B) + C$. + + \paragraph{(ii)}% + \label{par:exercise-4.15b-ii} + + Let $x \in (A + B) + C$. + Then $x$ is in $A + B$ or in $C$, but not both. + There are two cases to consider: + + \subparagraph{Case 1}% + + Suppose $x \in A + B$ and $x \not\in C$. + Then (a) $x \in A$ or (b) $x \in B$ but not both. + Suppose (a) was true. + That is, $x \in A$ and $x \not\in B$. + Since $x$ is not a member of $B$ nor $C$, $x \not\in B + C$. + Since $x$ is not a member of $B + C$ but is a member of $A$, + $x \in A + (B + C)$. + Now Suppose (b) was true. + That is, $x \not\in A$ and $x \in B$. + Since $x$ is a member of $B$ and not of $C$, then $x \in B + C$. + Since $x$ is a member of $B + C$ and not of $A$, $x \in A + (B + C)$. + + \subparagraph{Case 2}% + + Suppose $x \not\in A + B$ and $x \in C$. + Then, by \nameref{sub:nonmembership-symmetric-difference}, + (a) $x \in A \cap B$ or (b) $x \not\in A \cup B$. + Suppose (a) was true. + That is, $x \in A \land x \in B$. + Since $x$ is a member of $B$ and $C$, $x \not\in B + C$. + Since $x$ is not a member of $B + C$ but is a member of $A$, + $x \in A + (B + C)$. + Now suppose (b) was true. + That is, $x \not\in A$ and $x \not\in B$. + Since $x$ is not a member of $B$ but is a member of $C$, $x \in B + C$. + Since $x$ is a member of $B + C$ but not of $A$, $x \in A + (B + C)$. + + \paragraph{Conclusion}% + + In both \nameref{par:exercise-4.15b-i} and \nameref{par:exercise-4.15b-ii}, + the subcases are exhaustive and prove the desired subset relation. + Therefore $A + (B + C) = (A + B) + C$. \end{proof} diff --git a/Bookshelf/Enderton/Set/Chapter_2.lean b/Bookshelf/Enderton/Set/Chapter_2.lean index f7f336f..2290fd4 100644 --- a/Bookshelf/Enderton/Set/Chapter_2.lean +++ b/Bookshelf/Enderton/Set/Chapter_2.lean @@ -1,7 +1,9 @@ import Mathlib.Data.Set.Basic import Mathlib.Data.Set.Lattice +import Mathlib.Tactic.LibrarySearch import Bookshelf.Enderton.Set.Chapter_1 +import Common.Logic.Basic import Common.Set.Basic /-! # Enderton.Chapter_2 @@ -198,7 +200,7 @@ theorem exercise_3_7b_ii apply Iff.intro · intro h by_contra nh - rw [not_or] at nh + rw [not_or_de_morgan] at nh have ⟨a, hA⟩ := Set.not_subset.mp nh.left have ⟨b, hB⟩ := Set.not_subset.mp nh.right rw [Set.ext_iff] at h @@ -281,4 +283,139 @@ theorem exercise_3_10 (ha : a ∈ B) rw [← hb, Set.mem_setOf_eq] exact h₂ +/-- ### Exercise 4.11 (i) + +Show that for any sets `A` and `B`, `A = (A ∩ B) ∪ (A - B)`. +-/ +theorem exercise_4_11_i {A B : Set α} + : A = (A ∩ B) ∪ (A \ B) := by + unfold Union.union Set.instUnionSet Set.union + unfold SDiff.sdiff Set.instSDiffSet Set.diff + unfold Inter.inter Set.instInterSet Set.inter + unfold Membership.mem Set.instMembershipSet Set.Mem setOf + simp only + suffices ∀ x, (A x ∧ (B x ∨ ¬B x)) = A x by + conv => rhs; ext x; rw [← and_or_left, this] + intro x + refine propext ?_ + apply Iff.intro + · intro hx + exact hx.left + · intro hx + exact ⟨hx, em (B x)⟩ + +/-- ### Exercise 4.11 (ii) + +Show that for any sets `A` and `B`, `A ∪ (B - A) = A ∪ B`. +-/ +theorem exercise_4_11_ii {A B : Set α} + : A ∪ (B \ A) = A ∪ B := by + unfold Union.union Set.instUnionSet Set.union + unfold SDiff.sdiff Set.instSDiffSet Set.diff + unfold Membership.mem Set.instMembershipSet Set.Mem setOf + simp only + suffices ∀ x, ((A x ∨ B x) ∧ (A x ∨ ¬A x)) = (A x ∨ B x) by + conv => lhs; ext x; rw [or_and_left, this x] + intro x + refine propext ?_ + apply Iff.intro + · intro hx + exact hx.left + · intro hx + exact ⟨hx, em (A x)⟩ + +section + +variable {A B C : Set ℕ} + +variable {hA : A = {1, 2, 3}} +variable {hB : B = {2, 3, 4}} +variable {hC : C = {3, 4, 5}} + +lemma right_diff_eq_insert_one_three : A \ (B \ C) = {1, 3} := by + rw [hA, hB, hC] + ext x + apply Iff.intro + · intro hx + unfold SDiff.sdiff Set.instSDiffSet Set.diff at hx + unfold Membership.mem Set.instMembershipSet Set.Mem setOf at hx + unfold insert Set.instInsertSet Set.insert setOf at hx + have ⟨ha, hb⟩ := hx + rw [not_and_de_morgan, not_or_de_morgan] at hb + simp only [Set.mem_singleton_iff, not_not] at hb + refine Or.elim ha Or.inl ?_ + intro hy + apply Or.elim hb + · intro hz + exact Or.elim hy (absurd · hz.left) Or.inr + · intro hz + refine Or.elim hz Or.inr ?_ + intro hz₁ + apply Or.elim hy <;> apply Or.elim hz₁ <;> + · intro hz₂ hz₃ + rw [hz₂] at hz₃ + simp at hz₃ + · intro hx + unfold SDiff.sdiff Set.instSDiffSet Set.diff + unfold Membership.mem Set.instMembershipSet Set.Mem setOf + unfold insert Set.instInsertSet Set.insert setOf + apply Or.elim hx + · intro hy + refine ⟨Or.inl hy, ?_⟩ + intro hz + rw [hy] at hz + unfold Membership.mem Set.instMembershipSet Set.Mem at hz + unfold singleton Set.instSingletonSet Set.singleton setOf at hz + simp only at hz + · intro hy + refine ⟨Or.inr (Or.inr hy), ?_⟩ + intro hz + have hzr := hz.right + rw [not_or_de_morgan] at hzr + exact absurd hy hzr.left + +lemma left_diff_eq_singleton_one : (A \ B) \ C = {1} := by + rw [hA, hB, hC] + ext x + apply Iff.intro + · intro hx + unfold SDiff.sdiff Set.instSDiffSet Set.diff at hx + unfold Membership.mem Set.instMembershipSet Set.Mem setOf at hx + unfold insert Set.instInsertSet Set.insert setOf at hx + have ⟨⟨ha, hb⟩, hc⟩ := hx + rw [not_or_de_morgan] at hb hc + apply Or.elim ha + · simp + · intro hy + apply Or.elim hy + · intro hz + exact absurd hz hb.left + · intro hz + exact absurd hz hc.left + · intro hx + refine ⟨⟨Or.inl hx, ?_⟩, ?_⟩ <;> + · intro hy + cases hy with + | inl y => rw [hx] at y; simp at y + | inr hz => cases hz with + | inl y => rw [hx] at y; simp at y + | inr y => rw [hx] at y; simp at y + +/-- ### Exercise 4.14 + +Show by example that for some sets `A`, `B`, and `C`, the set `A - (B - C)` is +different from `(A - B) - C`. +-/ +theorem exercise_4_14 : A \ (B \ C) ≠ (A \ B) \ C := by + rw [ + @right_diff_eq_insert_one_three A B C hA hB hC, + @left_diff_eq_singleton_one A B C hA hB hC + ] + intro h + rw [Set.ext_iff] at h + have := h 3 + simp at this + +end + end Enderton.Set.Chapter_2 \ No newline at end of file diff --git a/Common.lean b/Common.lean index 0fe560b..2d5b4b8 100644 --- a/Common.lean +++ b/Common.lean @@ -1,5 +1,6 @@ import Common.Combinator import Common.Finset import Common.List +import Common.Logic import Common.Real import Common.Set \ No newline at end of file diff --git a/Common/Logic.lean b/Common/Logic.lean new file mode 100644 index 0000000..520ecd8 --- /dev/null +++ b/Common/Logic.lean @@ -0,0 +1 @@ +import Common.Logic.Basic \ No newline at end of file diff --git a/Common/Logic/Basic.lean b/Common/Logic/Basic.lean new file mode 100644 index 0000000..a385dd8 --- /dev/null +++ b/Common/Logic/Basic.lean @@ -0,0 +1,18 @@ +import Mathlib.Logic.Basic +import Mathlib.Tactic.Tauto + +/-! # Common.Logic.Basic + +Additional theorems and definitions related to basic logic. +-/ + +/-- +The de Morgan law that distributes negation across a conjunction. +-/ +theorem not_and_de_morgan {a b : Prop} : (¬(a ∧ b)) ↔ (¬ a ∨ ¬ b) := by + tauto + +/-- +Renaming of `not_or` to indicate its relationship to de Morgan's laws. +-/ +theorem not_or_de_morgan : ¬(p ∨ q) ↔ ¬p ∧ ¬q := not_or \ No newline at end of file diff --git a/Common/Set/Basic.lean b/Common/Set/Basic.lean index dd37419..a8e3ffd 100644 --- a/Common/Set/Basic.lean +++ b/Common/Set/Basic.lean @@ -1,5 +1,7 @@ import Mathlib.Data.Set.Basic +import Common.Logic.Basic + /-! # Common.Set.Basic Additional theorems and definitions useful in the context of `Set`s. @@ -87,4 +89,57 @@ theorem mem_mem_imp_pair_subset {x y : α} · intro hy' rwa [hy'] +/-! ## Symmetric Difference -/ + +/-- +`x` is a member of the `symmDiff` of `A` and `B` **iff** `x ∈ A ∧ x ∉ B` or +`x ∉ A ∧ x ∈ B`. +-/ +theorem mem_symm_diff_iff_exclusive_mem {A B : Set α} + : x ∈ (A ∆ B) ↔ (x ∈ A ∧ x ∉ B) ∨ (x ∉ A ∧ x ∈ B) := by + unfold symmDiff + apply Iff.intro + · intro hx + simp at hx + conv => arg 2; rw [and_comm] + exact hx + · intro hx + simp + conv => arg 2; rw [and_comm] + exact hx + +/-- +`x` is not a member of the `symmDiff` of `A` and `B` **iff** `x ∈ A ∩ B` or +`x ∉ A ∪ B`. + +This is the contraposition of `mem_symm_diff_iff_exclusive_mem`. +-/ +theorem not_mem_symm_diff_inter_or_not_union {A B : Set α} + : x ∉ (A ∆ B) ↔ (x ∈ A ∩ B) ∨ (x ∉ A ∪ B) := by + unfold symmDiff + simp + rw [ + not_or_de_morgan, + not_and_de_morgan, not_and_de_morgan, + not_not, not_not, + not_or_de_morgan + ] + apply Iff.intro + · intro hx + apply Or.elim hx.left + · intro nA + exact Or.elim hx.right + (fun nB => Or.inr ⟨nA, nB⟩) + (fun hA => absurd hA nA) + · intro hB + apply Or.elim hx.right + (fun nB => absurd hB nB) + (fun hA => Or.inl ⟨hA, hB⟩) + · intro hx + apply Or.elim hx + · intro hy + exact ⟨Or.inr hy.right, Or.inr hy.left⟩ + · intro hy + exact ⟨Or.inl hy.left, Or.inl hy.right⟩ + end Set \ No newline at end of file