Apostol Chapter 1.11, exercise 4c.

finite-set-exercises
Joshua Potter 2023-05-08 20:30:59 -06:00
parent ed89078e76
commit 0a8aa466c6
2 changed files with 31 additions and 4 deletions

View File

@ -49,7 +49,29 @@ theorem exercise_4b_2 (x : ) (h : ∃ n : , x ∈ Set.Ioo ↑n (↑n + (1
-/ -/
theorem exercise_4c (x y : ) theorem exercise_4c (x y : )
: ⌊x + y⌋ = ⌊x⌋ + ⌊y⌋ ⌊x + y⌋ = ⌊x⌋ + ⌊y⌋ + 1 := by : ⌊x + y⌋ = ⌊x⌋ + ⌊y⌋ ⌊x + y⌋ = ⌊x⌋ + ⌊y⌋ + 1 := by
sorry have hx : x = Int.floor x + Int.fract x := Eq.symm (add_eq_of_eq_sub' rfl)
have hy : y = Int.floor y + Int.fract y := Eq.symm (add_eq_of_eq_sub' rfl)
by_cases Int.fract x + Int.fract y < 1
· refine Or.inl ?_
rw [Int.floor_eq_iff]
simp only [Int.cast_add]
apply And.intro
· exact add_le_add (Int.floor_le x) (Int.floor_le y)
· conv => lhs; rw [hx, hy, add_add_add_comm]; arg 1; rw [add_comm]
rwa [add_comm, ← add_assoc, ← sub_lt_iff_lt_add', ← sub_sub, add_sub_cancel, add_sub_cancel]
· refine Or.inr ?_
rw [Int.floor_eq_iff]
simp only [Int.cast_add, Int.cast_one]
have h := le_of_not_lt h
apply And.intro
· conv => lhs; rw [← add_rotate]
conv => rhs; rw [hx, hy, add_add_add_comm]; arg 1; rw [add_comm]
rwa [← sub_le_iff_le_add', ← sub_sub, add_sub_cancel, add_sub_cancel]
· conv => lhs; rw [hx, hy, add_add_add_comm]; arg 1; rw [add_comm]
conv => lhs; rw [add_comm, ← add_assoc]
conv => rhs; rw [add_assoc]
rw [← sub_lt_iff_lt_add', ← sub_sub, add_sub_cancel, add_sub_cancel]
exact add_lt_add (Int.fract_lt_one x) (Int.fract_lt_one y)
/-- ### Exercise 4d /-- ### Exercise 4d
@ -73,7 +95,7 @@ The formulas in Exercises 4(d) and 4(e) suggest a generalization for `⌊nx⌋`.
State and prove such a generalization. State and prove such a generalization.
-/ -/
theorem exercise_5 (n : ) (x : ) theorem exercise_5 (n : ) (x : )
: ⌊n * x⌋ = 10 := by : True := by
sorry sorry
/-- ### Exercise 7b /-- ### Exercise 7b

View File

@ -39,7 +39,12 @@ $\floor{-x} =
\begin{proof} \begin{proof}
\link{exercise\_4b} \ % Force space prior to *Proof.*
\begin{enumerate}[(a)]
\item \link{exercise\_4b\_1}
\item \link{exercise\_4b\_2}
\end{enumerate}
\end{proof} \end{proof}
@ -136,7 +141,7 @@ Now apply Exercises 4(a) and (b) to the bracket on the right.
\begin{proof} \begin{proof}
TODO \link{exercise\_7b}
\end{proof} \end{proof}