Enderton (logic). Continue Exercises 1.1.
parent
456303f1dc
commit
94550ab43e
|
@ -602,7 +602,7 @@
|
||||||
|
|
||||||
\end{proof}
|
\end{proof}
|
||||||
|
|
||||||
\subsection{\pending{Exercise 1.1.5}}%
|
\subsection{\verified{Exercise 1.1.5}}%
|
||||||
\hyperlabel{sub:exercise-1.1.5}
|
\hyperlabel{sub:exercise-1.1.5}
|
||||||
|
|
||||||
Suppose that $\alpha$ is a wff not containing the negation symbol $\neg$.
|
Suppose that $\alpha$ is a wff not containing the negation symbol $\neg$.
|
||||||
|
@ -675,7 +675,7 @@
|
||||||
|
|
||||||
\end{proof}
|
\end{proof}
|
||||||
|
|
||||||
\subsubsection{\pending{Exercise 1.1.5b}}%
|
\subsubsection{\verified{Exercise 1.1.5b}}%
|
||||||
\hyperlabel{ssub:exercise-1.1.5-b}
|
\hyperlabel{ssub:exercise-1.1.5-b}
|
||||||
|
|
||||||
Show that more than a quarter of the symbols are sentence symbols.
|
Show that more than a quarter of the symbols are sentence symbols.
|
||||||
|
|
|
@ -2,6 +2,8 @@ import Common.Logic.Basic
|
||||||
import Common.Nat.Basic
|
import Common.Nat.Basic
|
||||||
import Mathlib.Algebra.Parity
|
import Mathlib.Algebra.Parity
|
||||||
import Mathlib.Data.Nat.Basic
|
import Mathlib.Data.Nat.Basic
|
||||||
|
import Mathlib.Data.Real.Basic
|
||||||
|
import Mathlib.Data.Setoid.Partition
|
||||||
import Mathlib.Tactic.NormNum
|
import Mathlib.Tactic.NormNum
|
||||||
import Mathlib.Tactic.Ring
|
import Mathlib.Tactic.Ring
|
||||||
|
|
||||||
|
@ -112,7 +114,22 @@ negation symbol is the only non-binary sentential connective.
|
||||||
-/
|
-/
|
||||||
lemma no_neg_sentential_count_eq_binary_count {φ : Wff} (h : ¬φ.hasNotSymbol)
|
lemma no_neg_sentential_count_eq_binary_count {φ : Wff} (h : ¬φ.hasNotSymbol)
|
||||||
: φ.sententialSymbolCount = φ.binarySymbolCount := by
|
: φ.sententialSymbolCount = φ.binarySymbolCount := by
|
||||||
sorry
|
induction φ with
|
||||||
|
| SS _ =>
|
||||||
|
unfold sententialSymbolCount binarySymbolCount
|
||||||
|
rfl
|
||||||
|
| Not _ _ =>
|
||||||
|
unfold hasNotSymbol at h
|
||||||
|
exfalso
|
||||||
|
exact h trivial
|
||||||
|
| And e₁ e₂ ih₁ ih₂
|
||||||
|
| Or e₁ e₂ ih₁ ih₂
|
||||||
|
| Cond e₁ e₂ ih₁ ih₂
|
||||||
|
| Iff e₁ e₂ ih₁ ih₂ =>
|
||||||
|
unfold hasNotSymbol at h
|
||||||
|
rw [not_or_de_morgan] at h
|
||||||
|
unfold sententialSymbolCount binarySymbolCount
|
||||||
|
rw [ih₁ h.left, ih₂ h.right]
|
||||||
|
|
||||||
/-- #### Parentheses Count
|
/-- #### Parentheses Count
|
||||||
|
|
||||||
|
@ -163,6 +180,14 @@ theorem length_eq_sum_symbol_count (φ : Wff)
|
||||||
|
|
||||||
end Wff
|
end Wff
|
||||||
|
|
||||||
|
/-! #### Exercise 1.1.2
|
||||||
|
|
||||||
|
Show that there are no wffs of length `2`, `3`, or `6`, but that any other
|
||||||
|
positive length is possible.
|
||||||
|
-/
|
||||||
|
|
||||||
|
section Exercise_1_1_2
|
||||||
|
|
||||||
/--
|
/--
|
||||||
An enumeration of values `m` and `n` can take on in equation `m + n = 3`.
|
An enumeration of values `m` and `n` can take on in equation `m + n = 3`.
|
||||||
-/
|
-/
|
||||||
|
@ -202,12 +227,6 @@ private lemma eq_3_by_cases (m n : ℕ) (h : m + n = 3)
|
||||||
right; right; right
|
right; right; right
|
||||||
exact ⟨m_eq_3, h⟩
|
exact ⟨m_eq_3, h⟩
|
||||||
|
|
||||||
/-! #### Exercise 1.1.2
|
|
||||||
|
|
||||||
Show that there are no wffs of length `2`, `3`, or `6`, but that any other
|
|
||||||
positive length is possible.
|
|
||||||
-/
|
|
||||||
|
|
||||||
theorem exercise_1_1_2_i (φ : Wff)
|
theorem exercise_1_1_2_i (φ : Wff)
|
||||||
: φ.length ≠ 2 ∧ φ.length ≠ 3 ∧ φ.length ≠ 6 := by
|
: φ.length ≠ 2 ∧ φ.length ≠ 3 ∧ φ.length ≠ 6 := by
|
||||||
induction φ with
|
induction φ with
|
||||||
|
@ -255,13 +274,28 @@ theorem exercise_1_1_2_i (φ : Wff)
|
||||||
intro h₃
|
intro h₃
|
||||||
exact absurd h₃.left ih₁.right.left
|
exact absurd h₃.left ih₁.right.left
|
||||||
|
|
||||||
theorem exercise_1_1_2_ii (n : ℕ) (h : n ∈ Set.univ \ {2, 3, 6})
|
private def recNot : ℕ → Wff → Wff
|
||||||
|
| 0, φ => φ
|
||||||
|
| n + 1, φ => Wff.Not (recNot n φ)
|
||||||
|
|
||||||
|
theorem exercise_1_1_2_ii (n : ℕ) (hn : n ≠ 2 ∧ n ≠ 3 ∧ n ≠ 6)
|
||||||
: ∃ φ : Wff, φ.length = n := by
|
: ∃ φ : Wff, φ.length = n := by
|
||||||
let φ₁ := Wff.SS 1
|
let φ₁ := Wff.SS 1
|
||||||
let φ₂ := Wff.And φ₁ (Wff.SS 2)
|
let φ₂ := Wff.And φ₁ (Wff.SS 2)
|
||||||
let φ₃ := Wff.And φ₂ (Wff.SS 3)
|
let φ₃ := Wff.And φ₂ (Wff.SS 3)
|
||||||
|
|
||||||
|
let S : Set (Set { n : ℕ // n ≠ 2 ∧ n ≠ 3 ∧ n ≠ 6 }) := {
|
||||||
|
{ m | ∃ n : ℕ, (recNot n φ₁).length = m.1 },
|
||||||
|
{ m | ∃ n : ℕ, (recNot n φ₂).length = m.1 },
|
||||||
|
{ m | ∃ n : ℕ, (recNot n φ₃).length = m.1 }
|
||||||
|
}
|
||||||
|
have hS : Setoid.IsPartition S := by
|
||||||
sorry
|
sorry
|
||||||
|
|
||||||
|
sorry
|
||||||
|
|
||||||
|
end Exercise_1_1_2
|
||||||
|
|
||||||
/-- #### Exercise 1.1.3
|
/-- #### Exercise 1.1.3
|
||||||
|
|
||||||
Let `α` be a wff; let `c` be the number of places at which binary connective
|
Let `α` be a wff; let `c` be the number of places at which binary connective
|
||||||
|
@ -330,10 +364,7 @@ than a quarter of the symbols are sentence symbols.
|
||||||
`k + 1`.
|
`k + 1`.
|
||||||
-/
|
-/
|
||||||
theorem exercise_1_1_5_b (α : Wff) (hα : ¬α.hasNotSymbol)
|
theorem exercise_1_1_5_b (α : Wff) (hα : ¬α.hasNotSymbol)
|
||||||
: α.sentenceSymbolCount > α.length / 4 := by
|
: α.sentenceSymbolCount > (Nat.cast α.length : ℝ) / 4 := by
|
||||||
have h₁ := α.sentenceSymbolCount
|
|
||||||
have h₂ := α.sententialSymbolCount
|
|
||||||
have h₃ := α.parenCount
|
|
||||||
rw [
|
rw [
|
||||||
α.length_eq_sum_symbol_count,
|
α.length_eq_sum_symbol_count,
|
||||||
Wff.paren_count_double_sentential_count α,
|
Wff.paren_count_double_sentential_count α,
|
||||||
|
@ -341,9 +372,22 @@ theorem exercise_1_1_5_b (α : Wff) (hα : ¬α.hasNotSymbol)
|
||||||
exercise_1_1_3 α
|
exercise_1_1_3 α
|
||||||
]
|
]
|
||||||
generalize Wff.binarySymbolCount α = k
|
generalize Wff.binarySymbolCount α = k
|
||||||
calc k + 1
|
simp only [
|
||||||
_ = 4 * (k + 1) / 4 := Eq.symm $ Nat.mul_div_cancel_left (k + 1) (by simp)
|
Nat.cast_add,
|
||||||
_ = (k + 4 + k + 2 * k) / 4 := by ring_nf
|
Nat.cast_one,
|
||||||
_ > (k + 1 + k + 2 * k) / 4 := by sorry
|
Nat.cast_mul,
|
||||||
|
Nat.cast_ofNat,
|
||||||
|
gt_iff_lt
|
||||||
|
]
|
||||||
|
ring_nf
|
||||||
|
simp only [
|
||||||
|
Int.ofNat_eq_coe,
|
||||||
|
Nat.cast_one,
|
||||||
|
Int.cast_one,
|
||||||
|
Nat.cast_ofNat,
|
||||||
|
one_div,
|
||||||
|
add_lt_add_iff_right
|
||||||
|
]
|
||||||
|
exact inv_lt_one (by norm_num)
|
||||||
|
|
||||||
end Enderton.Logic.Chapter_1
|
end Enderton.Logic.Chapter_1
|
||||||
|
|
Loading…
Reference in New Issue