Update for use with latest Mathlib version.

pigeonhole-redux
Joshua Potter 2023-11-08 00:26:03 -07:00
parent 05639fd07e
commit 4f371ac9b8
5 changed files with 15 additions and 16 deletions

View File

@ -118,7 +118,7 @@ Derive the result analytically as follows: By changing the index of summation,
note that `Σ_{n=1}^{b-1} ⌊na / b⌋ = Σ_{n=1}^{b-1} ⌊a(b - n) / b⌋`. Now apply note that `Σ_{n=1}^{b-1} ⌊na / b⌋ = Σ_{n=1}^{b-1} ⌊a(b - n) / b⌋`. Now apply
Exercises 4(a) and (b) to the bracket on the right. Exercises 4(a) and (b) to the bracket on the right.
-/ -/
theorem exercise_7b (ha : a > 0) (hb : b > 0) (hp : Nat.coprime a b) theorem exercise_7b (ha : a > 0) (hb : b > 0) (hp : Nat.Coprime a b)
: ∑ n in (Finset.range b).filter (· > 0), ⌊n * ((a : ) : ) / b⌋ = : ∑ n in (Finset.range b).filter (· > 0), ⌊n * ((a : ) : ) / b⌋ =
((a - 1) * (b - 1)) / 2 := by ((a - 1) * (b - 1)) / 2 := by
sorry sorry

View File

@ -475,7 +475,7 @@ theorem exercise_1_2_2b_iii {k : } (h : Odd k)
have ⟨r, hr⟩ := h have ⟨r, hr⟩ := h
refine ⟨r, hr, ?_⟩ refine ⟨r, hr, ?_⟩
by_contra nr by_contra nr
have : r = 0 := Nat.eq_zero_of_nonpos r nr have : r = 0 := Nat.eq_zero_of_not_pos nr
rw [this] at hr rw [this] at hr
simp only [mul_zero, zero_add] at hr simp only [mul_zero, zero_add] at hr
exact absurd hr hk exact absurd hr hk

View File

@ -171,8 +171,7 @@ theorem emptyset_identity_ii (A : Set α)
: A ∩ ∅ = ∅ := calc A ∩ ∅ : A ∩ ∅ = ∅ := calc A ∩ ∅
_ = { x | x ∈ A ∧ x ∈ ∅ } := rfl _ = { x | x ∈ A ∧ x ∈ ∅ } := rfl
_ = { x | x ∈ A ∧ False } := rfl _ = { x | x ∈ A ∧ False } := rfl
_ = { x | False } := by simp _ = ∅ := by simp
_ = ∅ := rfl
#check Set.inter_empty #check Set.inter_empty
@ -181,8 +180,7 @@ theorem emptyset_identity_iii (A C : Set α)
_ = { x | x ∈ A ∧ x ∈ C \ A } := rfl _ = { x | x ∈ A ∧ x ∈ C \ A } := rfl
_ = { x | x ∈ A ∧ (x ∈ C ∧ x ∉ A) } := rfl _ = { x | x ∈ A ∧ (x ∈ C ∧ x ∉ A) } := rfl
_ = { x | x ∈ C ∧ False } := by simp _ = { x | x ∈ C ∧ False } := by simp
_ = { x | False } := by simp _ = ∅ := by simp
_ = ∅ := rfl
#check Set.inter_diff_self #check Set.inter_diff_self
@ -636,7 +634,7 @@ lemma left_diff_eq_singleton_one : (A \ B) \ C = {1} := by
have ⟨⟨ha, hb⟩, hc⟩ := hx have ⟨⟨ha, hb⟩, hc⟩ := hx
rw [not_or_de_morgan] at hb hc rw [not_or_de_morgan] at hb hc
apply Or.elim ha apply Or.elim ha
· simp · simp
· intro hy · intro hy
apply Or.elim hy apply Or.elim hy
· intro hz · intro hz
@ -792,11 +790,11 @@ theorem exercise_2_17_ii {A B : Set α} (h : A \ B = ∅)
theorem exercise_2_17_iii {A B : Set α} (h : A B = B) theorem exercise_2_17_iii {A B : Set α} (h : A B = B)
: A ∩ B = A := by : A ∩ B = A := by
suffices A ⊆ B from Set.inter_eq_left_iff_subset.mpr this suffices A ⊆ B from Set.inter_eq_left.mpr this
exact Set.union_eq_right_iff_subset.mp h exact Set.union_eq_right.mp h
theorem exercise_2_17_iv {A B : Set α} (h : A ∩ B = A) theorem exercise_2_17_iv {A B : Set α} (h : A ∩ B = A)
: A ⊆ B := Set.inter_eq_left_iff_subset.mp h : A ⊆ B := Set.inter_eq_left.mp h
/-- #### Exercise 2.19 /-- #### Exercise 2.19
@ -954,7 +952,7 @@ theorem exercise_2_24b {𝓐 : Set (Set α)}
/-- #### Exercise 2.25 /-- #### Exercise 2.25
Is `A ( 𝓑)` always the same as ` { A X | X ∈ 𝓑 }`? If not, then under Is `A ( 𝓑)` always the same as ` { A X | X ∈ 𝓑 }`? If not, then under
what conditions does equality hold? what conditions does equality hold?
-/ -/
theorem exercise_2_25 {A : Set α} (𝓑 : Set (Set α)) theorem exercise_2_25 {A : Set α} (𝓑 : Set (Set α))
: (A (⋃₀ 𝓑) = ⋃₀ { A X | X ∈ 𝓑 }) ↔ (A = ∅ Set.Nonempty 𝓑) := by : (A (⋃₀ 𝓑) = ⋃₀ { A X | X ∈ 𝓑 }) ↔ (A = ∅ Set.Nonempty 𝓑) := by
@ -995,4 +993,4 @@ theorem exercise_2_25 {A : Set α} (𝓑 : Set (Set α))
_ = { x | ∃ t, t ∈ { y | ∃ X, X ∈ 𝓑 ∧ A X = y } ∧ x ∈ t } := by simp _ = { x | ∃ t, t ∈ { y | ∃ X, X ∈ 𝓑 ∧ A X = y } ∧ x ∈ t } := by simp
_ = ⋃₀ { A X | X ∈ 𝓑 } := rfl _ = ⋃₀ { A X | X ∈ 𝓑 } := rfl
end Enderton.Set.Chapter_2 end Enderton.Set.Chapter_2

View File

@ -1,4 +1,5 @@
import Mathlib.Logic.Basic import Mathlib.Logic.Basic
import Mathlib.Data.Set.Basic
import Mathlib.Tactic.Tauto import Mathlib.Tactic.Tauto
/-! # Common.Logic.Basic /-! # Common.Logic.Basic
@ -39,4 +40,4 @@ theorem forall_mem_comm {X : Set α} {Y : Set β} (p : α → β → Prop)
· intro h v hv u hu · intro h v hv u hu
exact h u hu v hv exact h u hu v hv
· intro h u hu v hv · intro h u hu v hv
exact h v hv u hu exact h v hv u hu

View File

@ -84,10 +84,10 @@ def shouldRender : ModuleMember → Bool
end ModuleMember end ModuleMember
inductive AnalyzeTask where inductive AnalyzeTask where
| loadAll (load : List Name) : AnalyzeTask | loadAll (load : Array Name) : AnalyzeTask
| loadAllLimitAnalysis (analyze : List Name) : AnalyzeTask | loadAllLimitAnalysis (analyze : Array Name) : AnalyzeTask
def AnalyzeTask.getLoad : AnalyzeTask → List Name def AnalyzeTask.getLoad : AnalyzeTask → Array Name
| loadAll load => load | loadAll load => load
| loadAllLimitAnalysis load => load | loadAllLimitAnalysis load => load