From f982012925171dfbf895c43b8a38ae5fedee88ca Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Thu, 17 Jun 2021 22:20:36 +0100 Subject: [PATCH 01/41] initial commit --- src/order/locally_finite.lean | 173 ++++++++++++++++++++++++++++++++++ 1 file changed, 173 insertions(+) create mode 100644 src/order/locally_finite.lean diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean new file mode 100644 index 0000000000000..34c2c491efc62 --- /dev/null +++ b/src/order/locally_finite.lean @@ -0,0 +1,173 @@ +/- +Copyright (c) 2021 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import data.finset.lattice +import data.finset.preimage +import data.finset.sort +import data.set.finite +import data.set.intervals.basic +import data.int.basic + +/-! +# Locally finite preorders +-/ + +universe u + +open set + +@[algebra] class locally_finite_order (α : Type u) [preorder α] := +(finset_Icc : α → α → finset α) +(coe_finset_Icc : ∀ x y : α, (finset_Icc x y : set α) = Icc x y) + +/- +@[algebra] class locally_finite_order (α : Type u) [preorder α] := +(list_Icc : α → α → list α) +(mem_Icc : ∀ {x y z}, z ∈ list_Icc x y ↔ z ∈ Icc x y) +-/ + +/- +@[algebra] class is_locally_finite_order (α : Type u) [preorder α] : Prop := +(Icc_finite : ∀ {x y : α}, (Icc x y).finite) +-/ + +variables {α : Type*} + +noncomputable def locally_finite_order_of_finite_Icc [preorder α] + (h : ∀ x y : α, (Icc x y).finite) : + locally_finite_order α := +{ finset_Icc := λ x y, (h x y).to_finset, + coe_finset_Icc := λ x y, (h x y).coe_to_finset } + +section +variables [preorder α] [locally_finite_order α] + +lemma Icc_finite (a b : α) : (Icc a b).finite := +begin + rw ←locally_finite_order.coe_finset_Icc, + exact (locally_finite_order.finset_Icc a b).finite_to_set, +end + +lemma Ico_finite (a b : α) : (Ico a b).finite := +(Icc_finite a b).subset Ico_subset_Icc_self + +lemma Ioc_finite (a b : α) : (Ioc a b).finite := +(Icc_finite a b).subset Ioc_subset_Icc_self + +lemma Ioo_finite (a b : α) : (Ioo a b).finite := +(Icc_finite a b).subset Ioo_subset_Icc_self + +end + +namespace finset +variables [partial_order α] [locally_finite_order α] [decidable_eq α] + +/-- `set.Icc a b` as a finset -/ +def Icc (a b : α) : finset α := +locally_finite_order.finset_Icc a b + +/-- `set.Icc a b` as a finset -/ +def Ico (a b : α) : finset α := +(Icc a b).erase b + +/-- `set.Icc a b` as a finset -/ +def Ioc (a b : α) : finset α := +(Icc a b).erase a + +/-- `set.Icc a b` as a finset -/ +def Ioo (a b : α) : finset α := +(Ioc a b).erase b + +@[simp] lemma coe_Icc (a b : α) : (Icc a b : set α) = set.Icc a b := +locally_finite_order.coe_finset_Icc a b + +@[simp] lemma coe_Ico (a b : α) : (Ico a b : set α) = set.Ico a b := +by rw [Ico, coe_erase, coe_Icc, Icc_diff_right] + +@[simp] lemma coe_Ioc (a b : α) : (Ioc a b : set α) = set.Ioc a b := +by rw [Ioc, coe_erase, coe_Icc, Icc_diff_left] + +@[simp] lemma coe_Ioo (a b : α) : (Ioo a b : set α) = set.Ioo a b := +by rw [Ioo, coe_erase, coe_Ioc, Ioc_diff_right] + +@[simp] lemma mem_Icc_iff {a b x : α} : x ∈ Icc a b ↔ x ∈ set.Icc a b := +by rw [←coe_Icc, mem_coe] + +@[simp] lemma mem_Ico_iff {a b x : α} : x ∈ Ico a b ↔ x ∈ set.Ico a b := +by rw [←coe_Ico, mem_coe] + +@[simp] lemma mem_Ioc_iff {a b x : α} : x ∈ Ioc a b ↔ x ∈ set.Ioc a b := +by rw [←coe_Ioc, mem_coe] + +@[simp] lemma mem_Ioo_iff {a b x : α} : x ∈ Ioo a b ↔ x ∈ set.Ioo a b := +by rw [←coe_Ioo, mem_coe] + +end finset + +namespace list +variables [partial_order α] [locally_finite_order α] [decidable_eq α] + [decidable_rel ((≤) : α → α → Prop)] + +/-/-- `set.Icc a b` as a finset -/ +def Icc (a b : α) : list α := +(finset.Icc a b).sort (≤)-/ + + + +end list + +lemma exists_min_greater [linear_order α] [locally_finite_order α] {x ub : α} (hx : x < ub) : + ∃ lub, x < lub ∧ ∀ y, x < y → lub ≤ y := +begin + have h : (finset.Ioc x ub).nonempty := ⟨ub, by rwa [finset.mem_Ioc_iff, right_mem_Ioc]⟩, + use finset.min' (finset.Ioc x ub) h, + split, + { + have := finset.min'_mem _ h, + simp * at *, + }, + rintro y hxy, + obtain hy | hy := le_or_lt y ub, + apply finset.min'_le, + simp * at *, + have := finset.min'_mem _ h, +end + +/-! ### Instances -/ + +variables {β : Type*} [partial_order β] [locally_finite_order β] [decidable_eq β] + +noncomputable def order_embedding.locally_finite_order [partial_order α] [decidable_eq α] + (f : α ↪o β) : + locally_finite_order α := +{ finset_Icc := λ a b, (finset.Icc (f a) (f b)).preimage f (f.to_embedding.injective.inj_on _), + coe_finset_Icc := λ a b, begin + ext, + simp only [finset.coe_Icc, finset.coe_preimage, iff_self, mem_preimage, mem_Icc, + order_embedding.le_iff_le], + end } + +@[priority 900] +instance fintype.locally_finite_order [preorder α] [fintype α] : locally_finite_order α := +{ finset_Icc := λ a b, (finite.of_fintype (Icc a b)).to_finset, + coe_finset_Icc := λ a b, finite.coe_to_finset _ } + +private def range : ℤ → ℕ → list ℤ +| z 0 := [] +| z (n + 1) := z :: range (z + 1) n + +instance int.locally_finite_order : locally_finite_order ℤ := +{ Icc_finite := λ a b, begin + let s := (range a (b - a).to_nat).to_finset, + apply (finset.finite_to_set s).subset, + rintro z hz, + apply set.finite_mem_finset, + refine set.finite.of_fintype _, +end } + +instance nat.locally_finite_order : locally_finite_order ℕ := +{ Icc_finite := λ x y, begin + +end } From 01c38cdbba8f375d7c800af3201d01a1ade71d30 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Fri, 18 Jun 2021 13:09:19 +0100 Subject: [PATCH 02/41] more of the module docstring, Ici, Ioi, Iic, Iio, cleanup --- src/order/locally_finite.lean | 253 +++++++++++++++++++++++++--------- 1 file changed, 188 insertions(+), 65 deletions(-) diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 34c2c491efc62..bb697e7f959d7 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -11,25 +11,90 @@ import data.set.intervals.basic import data.int.basic /-! -# Locally finite preorders +# Locally finite orders + +This file defines locally finite orders. + +A locally finite order is an order for which all bounded intervals are finite. This allows to make +sense of `Icc`/`Ico`/`Ioc`/`Ioo` as lists, multisets, or finsets. +Further, if the order is bounded above (resp. below), then we can also make sense of `Ici`/`Ioi` +(resp. `Iic`/`Iio`). + +## Main declarations + +In a `locally_finite_order`, +* `list`/`multiset`/`finset` dot `Icc`/`Ico`/`Ioc`/`Ioo`: Closed/open-closed/open interval as a + list/multiset/finset. +* `list`/`multiset`/`finset` dot `Ici`/`Ioi`: Closed/open-infinite interval as a + list/multiset/finset if the order is also an `order_top`. +* `list`/`multiset`/`finset` dot `Iic`/`Iio`: Infinite-closed/open interval as a + list/multiset/finset if the order is also an `order_bot`. + +## Instances + +We provide a `locally_finite_order` instance for +* ℕ, see `nat.locally_finite_order` +* ℤ, see `int.locally_finite_order` +* any fintype (but it is noncomputable), see `fintype.locally_finite_order` + +## TODO + +Once we have the `has_enum` typeclass (any non-top element has a least greater element or any +non-bot element has a greatest lesser element), we can provide an `has_enum` typeclass using + +lemma exists_min_greater [linear_order α] [locally_finite_order α] {x ub : α} (hx : x < ub) : + ∃ lub, x < lub ∧ ∀ y, x < y → lub ≤ y := +begin + have h : (finset.Ioc x ub).nonempty := ⟨ub, finset.mem_Ioc_iff.2 ⟨hx, le_rfl⟩⟩, + use finset.min' (finset.Ioc x ub) h, + split, + { + have := finset.min'_mem _ h, + simp * at *, + }, + rintro y hxy, + obtain hy | hy := le_total y ub, + apply finset.min'_le, + simp * at *, + exact (finset.min'_le _ _ (finset.mem_Ioc_iff.2 ⟨hx, le_rfl⟩)).trans hy, +end + +Note that the converse is not true. Consider `{0} ∪ {2^z | z : ℤ}`. Any element is either ⊥ or has +a greatest lesser element, so it `has_enum`, but it's not locally finite as `Icc 0 1` is infinite. + +/- more `has_enum` gibberish +def nth [linear_order α] [locally_finite_order α] {ub : α → option α} + (hub : ∀ a b, ub a = some b → a < b) (a : α) : ℕ → option α +| 0 := some a +| (n + 1) := match nth n with + | option.none := none + | (option.some a) := match ub a with + | option.none := none + | (option.some b) := classical.some (exists_min_greater (hub a b begin sorry end)) + +def nth' [linear_order α] [no_top_order α] [locally_finite_order α] (a : α) : ℕ → α +| 0 := a +| (n + 1) := classical.some (exists_min_greater (classical.some_spec (no_top a))) -/ -universe u +From `linear_order α`, `no_top_order α`, `locally_finite_order α`, we can also define an order +isomorphism `α ≃ ℕ` or `α ≃ ℤ` depending on whether we have `order_bot α` or `no_bot_order α`. +-/ open set -@[algebra] class locally_finite_order (α : Type u) [preorder α] := +class locally_finite_order (α : Type*) [preorder α] := (finset_Icc : α → α → finset α) (coe_finset_Icc : ∀ x y : α, (finset_Icc x y : set α) = Icc x y) /- -@[algebra] class locally_finite_order (α : Type u) [preorder α] := +class locally_finite_order (α : Type u) [preorder α] := (list_Icc : α → α → list α) (mem_Icc : ∀ {x y z}, z ∈ list_Icc x y ↔ z ∈ Icc x y) -/ /- -@[algebra] class is_locally_finite_order (α : Type u) [preorder α] : Prop := +class is_locally_finite_order (α : Type u) [preorder α] : Prop := (Icc_finite : ∀ {x y : α}, (Icc x y).finite) -/ @@ -41,7 +106,7 @@ noncomputable def locally_finite_order_of_finite_Icc [preorder α] { finset_Icc := λ x y, (h x y).to_finset, coe_finset_Icc := λ x y, (h x y).coe_to_finset } -section +section preorder variables [preorder α] [locally_finite_order α] lemma Icc_finite (a b : α) : (Icc a b).finite := @@ -59,24 +124,25 @@ lemma Ioc_finite (a b : α) : (Ioc a b).finite := lemma Ioo_finite (a b : α) : (Ioo a b).finite := (Icc_finite a b).subset Ioo_subset_Icc_self -end +end preorder namespace finset -variables [partial_order α] [locally_finite_order α] [decidable_eq α] +section partial_order +variables [decidable_eq α] [partial_order α] [locally_finite_order α] /-- `set.Icc a b` as a finset -/ def Icc (a b : α) : finset α := locally_finite_order.finset_Icc a b -/-- `set.Icc a b` as a finset -/ +/-- `set.Ico a b` as a finset -/ def Ico (a b : α) : finset α := (Icc a b).erase b -/-- `set.Icc a b` as a finset -/ +/-- `set.Ioc a b` as a finset -/ def Ioc (a b : α) : finset α := (Icc a b).erase a -/-- `set.Icc a b` as a finset -/ +/-- `set.Ioo a b` as a finset -/ def Ioo (a b : α) : finset α := (Ioc a b).erase b @@ -92,82 +158,139 @@ by rw [Ioc, coe_erase, coe_Icc, Icc_diff_left] @[simp] lemma coe_Ioo (a b : α) : (Ioo a b : set α) = set.Ioo a b := by rw [Ioo, coe_erase, coe_Ioc, Ioc_diff_right] -@[simp] lemma mem_Icc_iff {a b x : α} : x ∈ Icc a b ↔ x ∈ set.Icc a b := -by rw [←coe_Icc, mem_coe] +@[simp] lemma mem_Icc_iff {a b x : α} : x ∈ Icc a b ↔ a ≤ x ∧ x ≤ b := +by rw [←set.mem_Icc, ←coe_Icc, mem_coe] -@[simp] lemma mem_Ico_iff {a b x : α} : x ∈ Ico a b ↔ x ∈ set.Ico a b := -by rw [←coe_Ico, mem_coe] +@[simp] lemma mem_Ico_iff {a b x : α} : x ∈ Ico a b ↔ a ≤ x ∧ x < b := +by rw [←set.mem_Ico, ←coe_Ico, mem_coe] + +@[simp] lemma mem_Ioc_iff {a b x : α} : x ∈ Ioc a b ↔ a < x ∧ x ≤ b := +by rw [←set.mem_Ioc, ←coe_Ioc, mem_coe] + +@[simp] lemma mem_Ioo_iff {a b x : α} : x ∈ Ioo a b ↔ a < x ∧ x < b := +by rw [←set.mem_Ioo, ←coe_Ioo, mem_coe] + +@[simp] lemma right_mem_Ioc {a b : α} : b ∈ Ioc a b ↔ a < b := +begin + simp, +end -@[simp] lemma mem_Ioc_iff {a b x : α} : x ∈ Ioc a b ↔ x ∈ set.Ioc a b := -by rw [←coe_Ioc, mem_coe] +end partial_order -@[simp] lemma mem_Ioo_iff {a b x : α} : x ∈ Ioo a b ↔ x ∈ set.Ioo a b := -by rw [←coe_Ioo, mem_coe] +section order_top +variables [decidable_eq α] [order_top α] [locally_finite_order α] + +/-- `set.Ici a b` as a finset -/ +def Ici (a : α) : finset α := +Icc a ⊤ + +/-- `set.Ioi a b` as a finset -/ +def Ioi (a : α) : finset α := +Ioc a ⊤ + +end order_top + +section order_bot +variables [decidable_eq α] [order_bot α] [locally_finite_order α] + +/-- `set.Iic a b` as a finset -/ +def Iic (a : α) : finset α := +Icc ⊥ a + +/-- `set.Iio a b` as a finset -/ +def Iio (a : α) : finset α := +Ico ⊥ a + +end order_bot end finset +/- namespace list -variables [partial_order α] [locally_finite_order α] [decidable_eq α] +variables [decidable_eq α] [partial_order α] [locally_finite_order α] [decidable_rel ((≤) : α → α → Prop)] -/-/-- `set.Icc a b` as a finset -/ +/-- `set.Icc a b` as a list -/ def Icc (a b : α) : list α := -(finset.Icc a b).sort (≤)-/ - - +(finset.Icc a b).sort (≤) end list +-/ -lemma exists_min_greater [linear_order α] [locally_finite_order α] {x ub : α} (hx : x < ub) : - ∃ lub, x < lub ∧ ∀ y, x < y → lub ≤ y := +/-! ### Instances -/ + +/- +def locally_finite_order.order_iso_ℕ [order_bot α] [no_top_order α] [locally_finite_order α] : + order_iso α ℕ := begin - have h : (finset.Ioc x ub).nonempty := ⟨ub, by rwa [finset.mem_Ioc_iff, right_mem_Ioc]⟩, - use finset.min' (finset.Ioc x ub) h, - split, - { - have := finset.min'_mem _ h, - simp * at *, - }, - rintro y hxy, - obtain hy | hy := le_or_lt y ub, - apply finset.min'_le, - simp * at *, - have := finset.min'_mem _ h, -end -/-! ### Instances -/ +end-/ + +private def range₂ [semiring α] [decidable_eq α] : α → ℕ → finset α +| a 0 := ∅ +| a (n + 1) := insert a (range₂ (a + 1) n) + +private lemma mem_range₂_ℤ (x : ℤ) : ∀ a n, x ∈ range₂ a n ↔ a ≤ x ∧ x < a + n +| a 0 := by simp only [range₂, finset.not_mem_empty, add_zero, imp_self, false_iff, + int.coe_nat_zero, not_and, not_lt] +| a (n + 1) := begin + have h : x = a ∨ x < a + (↑n + 1) ↔ x < a + (↑n + 1), + { rw or_iff_right_iff_imp, + rintro rfl, + exact int.lt_add_succ _ _ }, + rw [range₂, finset.mem_insert, le_iff_eq_or_lt, mem_range₂_ℤ _ n, int.coe_nat_succ, add_assoc, + add_comm (1 : ℤ), @eq_comm _ a, or_and_distrib_left, h], + refl, +end -variables {β : Type*} [partial_order β] [locally_finite_order β] [decidable_eq β] +private lemma mem_range₂_ℕ (x : ℕ) : ∀ a n, x ∈ range₂ a n ↔ a ≤ x ∧ x < a + n +| a 0 := by simp only [range₂, finset.not_mem_empty, add_zero, imp_self, false_iff, + int.coe_nat_zero, not_and, not_lt] +| a (n + 1) := begin + have h : x = a ∨ x < a + (n + 1) ↔ x < a + (n + 1), + { rw or_iff_right_iff_imp, + rintro rfl, + exact (lt_add_iff_pos_right _).2 nat.succ_pos' }, + rw [range₂, finset.mem_insert, le_iff_eq_or_lt, mem_range₂_ℕ _ n, add_assoc, + add_comm 1, @eq_comm _ a, or_and_distrib_left, h], + refl, + end -noncomputable def order_embedding.locally_finite_order [partial_order α] [decidable_eq α] - (f : α ↪o β) : - locally_finite_order α := -{ finset_Icc := λ a b, (finset.Icc (f a) (f b)).preimage f (f.to_embedding.injective.inj_on _), +instance int.locally_finite_order : locally_finite_order ℤ := +{ finset_Icc := λ a b, range₂ a (b + 1 - a).to_nat, coe_finset_Icc := λ a b, begin ext, - simp only [finset.coe_Icc, finset.coe_preimage, iff_self, mem_preimage, mem_Icc, - order_embedding.le_iff_le], + rw [finset.mem_coe, mem_range₂_ℤ, set.mem_Icc], + cases le_or_lt a b, + { rw [int.to_nat_of_nonneg (sub_nonneg.2 (h.trans (lt_add_one _).le)), ←add_sub_assoc, + add_sub_cancel', int.lt_add_one_iff] }, + rw [int.to_nat_of_nonpos (sub_nonpos.2 h), int.coe_nat_zero, add_zero], + exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)), end } +instance nat.locally_finite_order : locally_finite_order ℕ := +{ finset_Icc := λ a b, range₂ a (b + 1 - a), + coe_finset_Icc := λ a b, begin + ext, + rw [finset.mem_coe, mem_range₂_ℕ, set.mem_Icc], + cases le_or_lt a b, + { rw [nat.add_sub_cancel' (nat.lt_succ_of_le h).le, nat.lt_succ_iff] }, + rw [nat.sub_eq_zero_iff_le.2 h, add_zero], + exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)), +end } + @[priority 900] -instance fintype.locally_finite_order [preorder α] [fintype α] : locally_finite_order α := +noncomputable instance fintype.locally_finite_order [preorder α] [fintype α] : + locally_finite_order α := { finset_Icc := λ a b, (finite.of_fintype (Icc a b)).to_finset, coe_finset_Icc := λ a b, finite.coe_to_finset _ } -private def range : ℤ → ℕ → list ℤ -| z 0 := [] -| z (n + 1) := z :: range (z + 1) n - -instance int.locally_finite_order : locally_finite_order ℤ := -{ Icc_finite := λ a b, begin - let s := (range a (b - a).to_nat).to_finset, - apply (finset.finite_to_set s).subset, - rintro z hz, - apply set.finite_mem_finset, - refine set.finite.of_fintype _, -end } - -instance nat.locally_finite_order : locally_finite_order ℕ := -{ Icc_finite := λ x y, begin - -end } +noncomputable def order_embedding.locally_finite_order {β : Type*} [partial_order β] + [locally_finite_order β] [decidable_eq β] [partial_order α] [decidable_eq α] (f : α ↪o β) : + locally_finite_order α := +{ finset_Icc := λ a b, (finset.Icc (f a) (f b)).preimage f (f.to_embedding.injective.inj_on _), + coe_finset_Icc := λ a b, begin + ext, + simp only [finset.coe_Icc, finset.coe_preimage, iff_self, mem_preimage, mem_Icc, + order_embedding.le_iff_le], + end } From 4303fa11e00f597c68bd4436054d07b7f7be49bb Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Fri, 18 Jun 2021 18:36:56 +0100 Subject: [PATCH 03/41] Iic, Iio, Ioc, Ioi --- src/order/locally_finite.lean | 100 ++++++++++++++++++++++------------ 1 file changed, 65 insertions(+), 35 deletions(-) diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index bb697e7f959d7..11b1516bd9625 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -17,8 +17,8 @@ This file defines locally finite orders. A locally finite order is an order for which all bounded intervals are finite. This allows to make sense of `Icc`/`Ico`/`Ioc`/`Ioo` as lists, multisets, or finsets. -Further, if the order is bounded above (resp. below), then we can also make sense of `Ici`/`Ioi` -(resp. `Iic`/`Iio`). +Further, if the order is bounded above (resp. below), then we can also make sense of the +"unbounded" intervals `Ici`/`Ioi` (resp. `Iic`/`Iio`). ## Main declarations @@ -180,26 +180,50 @@ end partial_order section order_top variables [decidable_eq α] [order_top α] [locally_finite_order α] -/-- `set.Ici a b` as a finset -/ +/-- `set.Ici a` as a finset -/ def Ici (a : α) : finset α := Icc a ⊤ -/-- `set.Ioi a b` as a finset -/ +/-- `set.Ioi a` as a finset -/ def Ioi (a : α) : finset α := Ioc a ⊤ +@[simp] lemma coe_Ici (a : α) : (Ici a : set α) = set.Ici a := +by rw [Ici, coe_Icc, Icc_top] + +@[simp] lemma coe_Ioi (a : α) : (Ioi a : set α) = set.Ioi a := +by rw [Ioi, coe_Ioc, Ioc_top] + +@[simp] lemma mem_Ici_iff {a x : α} : x ∈ Ici a ↔ a ≤ x := +by rw [←set.mem_Ici, ←coe_Ici, mem_coe] + +@[simp] lemma mem_Ioi_iff {a x : α} : x ∈ Ioi a ↔ a < x := +by rw [←set.mem_Ioi, ←coe_Ioi, mem_coe] + end order_top section order_bot variables [decidable_eq α] [order_bot α] [locally_finite_order α] -/-- `set.Iic a b` as a finset -/ -def Iic (a : α) : finset α := -Icc ⊥ a +/-- `set.Iic b` as a finset -/ +def Iic (b : α) : finset α := +Icc ⊥ b + +/-- `set.Iio b` as a finset -/ +def Iio (b : α) : finset α := +Ico ⊥ b + +@[simp] lemma coe_Iic (b : α) : (Iic b : set α) = set.Iic b := +by rw [Iic, coe_Icc, Icc_bot] + +@[simp] lemma coe_Iio (b : α) : (Iio b : set α) = set.Iio b := +by rw [Iio, coe_Ico, Ico_bot] -/-- `set.Iio a b` as a finset -/ -def Iio (a : α) : finset α := -Ico ⊥ a +@[simp] lemma mem_Iic_iff {b x : α} : x ∈ Iic b ↔ x ≤ b := +by rw [←set.mem_Iic, ←coe_Iic, mem_coe] + +@[simp] lemma mem_Iio_iff {b x : α} : x ∈ Iio b ↔ x < b := +by rw [←set.mem_Iio, ←coe_Iio, mem_coe] end order_bot @@ -243,24 +267,41 @@ private lemma mem_range₂_ℤ (x : ℤ) : ∀ a n, x ∈ range₂ a n ↔ a ≤ refl, end -private lemma mem_range₂_ℕ (x : ℕ) : ∀ a n, x ∈ range₂ a n ↔ a ≤ x ∧ x < a + n -| a 0 := by simp only [range₂, finset.not_mem_empty, add_zero, imp_self, false_iff, - int.coe_nat_zero, not_and, not_lt] -| a (n + 1) := begin - have h : x = a ∨ x < a + (n + 1) ↔ x < a + (n + 1), - { rw or_iff_right_iff_imp, - rintro rfl, - exact (lt_add_iff_pos_right _).2 nat.succ_pos' }, - rw [range₂, finset.mem_insert, le_iff_eq_or_lt, mem_range₂_ℕ _ n, add_assoc, - add_comm 1, @eq_comm _ a, or_and_distrib_left, h], - refl, - end +instance nat.locally_finite_order : locally_finite_order ℕ := +{ finset_Icc := λ a b, (list.range' a (b + 1 - a)).to_finset, + coe_finset_Icc := λ a b, begin + ext, + rw [finset.mem_coe, list.mem_to_finset, list.mem_range', set.mem_Icc], + cases le_or_lt a b, + { rw [nat.add_sub_cancel' (nat.lt_succ_of_le h).le, nat.lt_succ_iff] }, + rw [nat.sub_eq_zero_iff_le.2 h, add_zero], + exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)), +end } instance int.locally_finite_order : locally_finite_order ℤ := -{ finset_Icc := λ a b, range₂ a (b + 1 - a).to_nat, +{ finset_Icc := λ a b, (finset.Iio (b + 1 - a).to_nat).map ⟨λ n, n + a, λ _, by simp only + [imp_self, forall_const, add_left_inj, int.coe_nat_inj']⟩, coe_finset_Icc := λ a b, begin ext, - rw [finset.mem_coe, mem_range₂_ℤ, set.mem_Icc], + rw [finset.mem_coe, finset.mem_map, set.mem_Icc, function.embedding.coe_fn_mk], + split, + { + rintro ⟨n, hn, hx⟩, + rw finset.mem_Iio_iff at hn, + rw ←hx, + sorry + }, + rintro h, + refine ⟨(x - a).to_nat, _, _⟩, + { rw [finset.mem_Iio_iff, ←int.coe_nat_lt, int.to_nat_sub_of_le _ _ h.1], + exact int.to_nat_lt_to_nat (sub_le_sub_right h.2 _) }, + rw sub_eq_add_neg, + rw int.to_nat_sub_of_le _ _ h.1, + simp, + cases le_or_lt a b, + { + rw int.to_nat_sub_of_le _ _ h, + }, cases le_or_lt a b, { rw [int.to_nat_of_nonneg (sub_nonneg.2 (h.trans (lt_add_one _).le)), ←add_sub_assoc, add_sub_cancel', int.lt_add_one_iff] }, @@ -268,17 +309,6 @@ instance int.locally_finite_order : locally_finite_order ℤ := exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)), end } -instance nat.locally_finite_order : locally_finite_order ℕ := -{ finset_Icc := λ a b, range₂ a (b + 1 - a), - coe_finset_Icc := λ a b, begin - ext, - rw [finset.mem_coe, mem_range₂_ℕ, set.mem_Icc], - cases le_or_lt a b, - { rw [nat.add_sub_cancel' (nat.lt_succ_of_le h).le, nat.lt_succ_iff] }, - rw [nat.sub_eq_zero_iff_le.2 h, add_zero], - exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)), -end } - @[priority 900] noncomputable instance fintype.locally_finite_order [preorder α] [fintype α] : locally_finite_order α := From 9e117ecf1355667eff789a898e714ddde73472a1 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sun, 20 Jun 2021 20:47:49 +0100 Subject: [PATCH 04/41] polish --- src/order/locally_finite.lean | 267 +++++++++++++++------------------- 1 file changed, 114 insertions(+), 153 deletions(-) diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 11b1516bd9625..e5010c7c34644 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -23,91 +23,62 @@ Further, if the order is bounded above (resp. below), then we can also make sens ## Main declarations In a `locally_finite_order`, -* `list`/`multiset`/`finset` dot `Icc`/`Ico`/`Ioc`/`Ioo`: Closed/open-closed/open interval as a - list/multiset/finset. -* `list`/`multiset`/`finset` dot `Ici`/`Ioi`: Closed/open-infinite interval as a - list/multiset/finset if the order is also an `order_top`. -* `list`/`multiset`/`finset` dot `Iic`/`Iio`: Infinite-closed/open interval as a - list/multiset/finset if the order is also an `order_bot`. +* `finset.` `Icc`/`Ico`/`Ioc`/`Ioo`: Closed/open-closed/open interval as a finset. +* `finset.` `Ici`/`Ioi`: Closed/open-infinite interval as a finset if the order is also an + `order_top`. +* `finset.` `Iic`/`Iio`: Infinite-closed/open interval as a finset if the order is also an + `order_bot`. ## Instances We provide a `locally_finite_order` instance for * ℕ, see `nat.locally_finite_order` * ℤ, see `int.locally_finite_order` +* a subtype of a locally finite order * any fintype (but it is noncomputable), see `fintype.locally_finite_order` ## TODO -Once we have the `has_enum` typeclass (any non-top element has a least greater element or any +TODO: All the API in `set.intervals.basic` could be copied here. But this is unfortunately labor- +and maintenance-intensive. + +TODO:nce we have the `has_enum` typeclass (any non-top element has a least greater element or any non-bot element has a greatest lesser element), we can provide an `has_enum` typeclass using +```lean lemma exists_min_greater [linear_order α] [locally_finite_order α] {x ub : α} (hx : x < ub) : ∃ lub, x < lub ∧ ∀ y, x < y → lub ≤ y := -begin +begin -- very non golfed have h : (finset.Ioc x ub).nonempty := ⟨ub, finset.mem_Ioc_iff.2 ⟨hx, le_rfl⟩⟩, use finset.min' (finset.Ioc x ub) h, split, - { - have := finset.min'_mem _ h, - simp * at *, - }, + { have := finset.min'_mem _ h, + simp * at * }, rintro y hxy, obtain hy | hy := le_total y ub, apply finset.min'_le, simp * at *, exact (finset.min'_le _ _ (finset.mem_Ioc_iff.2 ⟨hx, le_rfl⟩)).trans hy, end +``` -Note that the converse is not true. Consider `{0} ∪ {2^z | z : ℤ}`. Any element is either ⊥ or has -a greatest lesser element, so it `has_enum`, but it's not locally finite as `Icc 0 1` is infinite. - -/- more `has_enum` gibberish -def nth [linear_order α] [locally_finite_order α] {ub : α → option α} - (hub : ∀ a b, ub a = some b → a < b) (a : α) : ℕ → option α -| 0 := some a -| (n + 1) := match nth n with - | option.none := none - | (option.some a) := match ub a with - | option.none := none - | (option.some b) := classical.some (exists_min_greater (hub a b begin sorry end)) - -def nth' [linear_order α] [no_top_order α] [locally_finite_order α] (a : α) : ℕ → α -| 0 := a -| (n + 1) := classical.some (exists_min_greater (classical.some_spec (no_top a))) --/ +Note that the converse is not true. Consider `{-2^z | z : ℤ} ∪ {2^z | z : ℤ}`. Any element has a +least greater element and a greatest lesser element, so it `has_enum` (both ways!), but it's not +locally finite as `Icc (-1) 1` is infinite. -From `linear_order α`, `no_top_order α`, `locally_finite_order α`, we can also define an order -isomorphism `α ≃ ℕ` or `α ≃ ℤ` depending on whether we have `order_bot α` or `no_bot_order α`. +TODO: From `linear_order α`, `no_top_order α`, `locally_finite_order α`, we can also define an +order isomorphism `α ≃ ℕ` or `α ≃ ℤ`, depending on whether we have `order_bot α` or +`no_bot_order α` and `nonempty α`. When `order_bot α`, we can match `a : α` to `(Iio a).card`. -/ -open set +open finset class locally_finite_order (α : Type*) [preorder α] := (finset_Icc : α → α → finset α) -(coe_finset_Icc : ∀ x y : α, (finset_Icc x y : set α) = Icc x y) - -/- -class locally_finite_order (α : Type u) [preorder α] := -(list_Icc : α → α → list α) -(mem_Icc : ∀ {x y z}, z ∈ list_Icc x y ↔ z ∈ Icc x y) --/ - -/- -class is_locally_finite_order (α : Type u) [preorder α] : Prop := -(Icc_finite : ∀ {x y : α}, (Icc x y).finite) --/ - -variables {α : Type*} - -noncomputable def locally_finite_order_of_finite_Icc [preorder α] - (h : ∀ x y : α, (Icc x y).finite) : - locally_finite_order α := -{ finset_Icc := λ x y, (h x y).to_finset, - coe_finset_Icc := λ x y, (h x y).coe_to_finset } +(coe_finset_Icc : ∀ a b : α, (finset_Icc a b : set α) = set.Icc a b) -section preorder -variables [preorder α] [locally_finite_order α] +namespace set +variables {α : Type*} [preorder α] [locally_finite_order α] lemma Icc_finite (a b : α) : (Icc a b).finite := begin @@ -124,42 +95,51 @@ lemma Ioc_finite (a b : α) : (Ioc a b).finite := lemma Ioo_finite (a b : α) : (Ioo a b).finite := (Icc_finite a b).subset Ioo_subset_Icc_self -end preorder +end set + +/-! ### Intervals as finsets -/ namespace finset -section partial_order -variables [decidable_eq α] [partial_order α] [locally_finite_order α] +section preorder +variables {α : Type*} [decidable_eq α] [preorder α] [locally_finite_order α] -/-- `set.Icc a b` as a finset -/ +/-- The finset of elements `x` such that `a ≤ x ∧ x ≤ b`. Basically `set.Icc a b` as a finset. -/ def Icc (a b : α) : finset α := locally_finite_order.finset_Icc a b -/-- `set.Ico a b` as a finset -/ +/-- The finset of elements `x` such that `a ≤ x ∧ x < b`. Basically `set.Ico a b` as a finset. -/ def Ico (a b : α) : finset α := (Icc a b).erase b -/-- `set.Ioc a b` as a finset -/ +/-- The finset of elements `x` such that `a < x ∧ x ≤ b`. Basically `set.Ioc a b` as a finset. -/ def Ioc (a b : α) : finset α := (Icc a b).erase a -/-- `set.Ioo a b` as a finset -/ +/-- The finset of elements `x` such that `a < x ∧ x < b`. Basically `set.Ioo a b` as a finset. -/ def Ioo (a b : α) : finset α := (Ioc a b).erase b +/- `Icc` is treated separately from `Ico`, `Ioc`, `Ioo` because it allows to relax `partial_order` +to `preorder` in simple constructions, eg `subtype.locally_finite_order`. -/ @[simp] lemma coe_Icc (a b : α) : (Icc a b : set α) = set.Icc a b := locally_finite_order.coe_finset_Icc a b +@[simp] lemma mem_Icc_iff {a b x : α} : x ∈ Icc a b ↔ a ≤ x ∧ x ≤ b := +by rw [←set.mem_Icc, ←coe_Icc, mem_coe] + +end preorder + +section partial_order +variables {α : Type*} [decidable_eq α] [partial_order α] [locally_finite_order α] + @[simp] lemma coe_Ico (a b : α) : (Ico a b : set α) = set.Ico a b := -by rw [Ico, coe_erase, coe_Icc, Icc_diff_right] +by rw [Ico, coe_erase, coe_Icc, set.Icc_diff_right] @[simp] lemma coe_Ioc (a b : α) : (Ioc a b : set α) = set.Ioc a b := -by rw [Ioc, coe_erase, coe_Icc, Icc_diff_left] +by rw [Ioc, coe_erase, coe_Icc, set.Icc_diff_left] @[simp] lemma coe_Ioo (a b : α) : (Ioo a b : set α) = set.Ioo a b := -by rw [Ioo, coe_erase, coe_Ioc, Ioc_diff_right] - -@[simp] lemma mem_Icc_iff {a b x : α} : x ∈ Icc a b ↔ a ≤ x ∧ x ≤ b := -by rw [←set.mem_Icc, ←coe_Icc, mem_coe] +by rw [Ioo, coe_erase, coe_Ioc, set.Ioc_diff_right] @[simp] lemma mem_Ico_iff {a b x : α} : x ∈ Ico a b ↔ a ≤ x ∧ x < b := by rw [←set.mem_Ico, ←coe_Ico, mem_coe] @@ -170,29 +150,24 @@ by rw [←set.mem_Ioc, ←coe_Ioc, mem_coe] @[simp] lemma mem_Ioo_iff {a b x : α} : x ∈ Ioo a b ↔ a < x ∧ x < b := by rw [←set.mem_Ioo, ←coe_Ioo, mem_coe] -@[simp] lemma right_mem_Ioc {a b : α} : b ∈ Ioc a b ↔ a < b := -begin - simp, -end - end partial_order section order_top -variables [decidable_eq α] [order_top α] [locally_finite_order α] +variables {α : Type*} [decidable_eq α] [order_top α] [locally_finite_order α] -/-- `set.Ici a` as a finset -/ +/-- The finset of elements `x` such that `a ≤ x`. Basically `set.Ici a` as a finset. -/ def Ici (a : α) : finset α := Icc a ⊤ -/-- `set.Ioi a` as a finset -/ +/-- The finset of elements `x` such that `a < x`. Basically `set.Ioi a` as a finset. -/ def Ioi (a : α) : finset α := Ioc a ⊤ @[simp] lemma coe_Ici (a : α) : (Ici a : set α) = set.Ici a := -by rw [Ici, coe_Icc, Icc_top] +by rw [Ici, coe_Icc, set.Icc_top] @[simp] lemma coe_Ioi (a : α) : (Ioi a : set α) = set.Ioi a := -by rw [Ioi, coe_Ioc, Ioc_top] +by rw [Ioi, coe_Ioc, set.Ioc_top] @[simp] lemma mem_Ici_iff {a x : α} : x ∈ Ici a ↔ a ≤ x := by rw [←set.mem_Ici, ←coe_Ici, mem_coe] @@ -203,21 +178,21 @@ by rw [←set.mem_Ioi, ←coe_Ioi, mem_coe] end order_top section order_bot -variables [decidable_eq α] [order_bot α] [locally_finite_order α] +variables {α : Type*} [decidable_eq α] [order_bot α] [locally_finite_order α] -/-- `set.Iic b` as a finset -/ +/-- The finset of elements `x` such that `x ≤ b`. Basically `set.Iic b` as a finset. -/ def Iic (b : α) : finset α := Icc ⊥ b -/-- `set.Iio b` as a finset -/ +/-- The finset of elements `x` such that `x < b`. Basically `set.Iio b` as a finset. -/ def Iio (b : α) : finset α := Ico ⊥ b @[simp] lemma coe_Iic (b : α) : (Iic b : set α) = set.Iic b := -by rw [Iic, coe_Icc, Icc_bot] +by rw [Iic, coe_Icc, set.Icc_bot] @[simp] lemma coe_Iio (b : α) : (Iio b : set α) = set.Iio b := -by rw [Iio, coe_Ico, Ico_bot] +by rw [Iio, coe_Ico, set.Ico_bot] @[simp] lemma mem_Iic_iff {b x : α} : x ∈ Iic b ↔ x ≤ b := by rw [←set.mem_Iic, ←coe_Iic, mem_coe] @@ -229,49 +204,64 @@ end order_bot end finset -/- -namespace list -variables [decidable_eq α] [partial_order α] [locally_finite_order α] - [decidable_rel ((≤) : α → α → Prop)] +open finset -/-- `set.Icc a b` as a list -/ -def Icc (a b : α) : list α := -(finset.Icc a b).sort (≤) +/-! ### Instances -/ -end list --/ +noncomputable def locally_finite_order_of_finite_Icc {α : Type*} [preorder α] + (h : ∀ x y : α, (set.Icc x y).finite) : + locally_finite_order α := +{ finset_Icc := λ x y, (h x y).to_finset, + coe_finset_Icc := λ x y, (h x y).coe_to_finset } -/-! ### Instances -/ +@[priority 900] +noncomputable instance fintype.locally_finite_order {α : Type*} [preorder α] [fintype α] : + locally_finite_order α := +{ finset_Icc := λ a b, (set.finite.of_fintype (set.Icc a b)).to_finset, + coe_finset_Icc := λ a b, set.finite.coe_to_finset _ } -/- -def locally_finite_order.order_iso_ℕ [order_bot α] [no_top_order α] [locally_finite_order α] : - order_iso α ℕ := -begin +-- Should this be called `order_embedding.locally_finite_order`? +/-- Given an order embedding `α ↪o β`, pulls back to `α` the `locally_finite_order` on `β`. -/ +noncomputable def locally_finite_order.lift {α β : Type*} [partial_order β] [locally_finite_order β] + [decidable_eq β] [partial_order α] [decidable_eq α] (f : α ↪o β) : + locally_finite_order α := +{ finset_Icc := λ a b, (Icc (f a) (f b)).preimage f (f.to_embedding.injective.inj_on _), + coe_finset_Icc := λ a b, begin + ext, + simp only [coe_Icc, coe_preimage, iff_self, set.mem_preimage, set.mem_Icc, + order_embedding.le_iff_le], + end } -end-/ - -private def range₂ [semiring α] [decidable_eq α] : α → ℕ → finset α -| a 0 := ∅ -| a (n + 1) := insert a (range₂ (a + 1) n) - -private lemma mem_range₂_ℤ (x : ℤ) : ∀ a n, x ∈ range₂ a n ↔ a ≤ x ∧ x < a + n -| a 0 := by simp only [range₂, finset.not_mem_empty, add_zero, imp_self, false_iff, - int.coe_nat_zero, not_and, not_lt] -| a (n + 1) := begin - have h : x = a ∨ x < a + (↑n + 1) ↔ x < a + (↑n + 1), - { rw or_iff_right_iff_imp, - rintro rfl, - exact int.lt_add_succ _ _ }, - rw [range₂, finset.mem_insert, le_iff_eq_or_lt, mem_range₂_ℤ _ n, int.coe_nat_succ, add_assoc, - add_comm (1 : ℤ), @eq_comm _ a, or_and_distrib_left, h], - refl, +instance subtype.locally_finite_order {α : Type*} [preorder α] [locally_finite_order α] + [decidable_eq α] {p : α → Prop} [decidable_pred p] : + locally_finite_order (subtype p) := +begin + haveI : preorder (subtype p) := by apply_instance, + haveI : decidable_eq (subtype p) := by apply_instance, + exact + { finset_Icc := λ a b, finset.subtype p (Icc a.val b.val), + coe_finset_Icc := λ a b, begin + ext, + rw [finset.subtype, coe_map, set.mem_image, set.mem_Icc], + dsimp, + split, + { rintro ⟨⟨y, hy⟩, -, h⟩, + rw [←h, ←subtype.coe_le_coe, ←subtype.coe_le_coe], + rw [mem_filter, mem_Icc_iff] at hy, + exact hy.1 }, + rintro hx, + refine ⟨⟨x, _⟩, mem_coe.2 (mem_attach _ _), _⟩, + { rw [mem_filter, mem_Icc_iff, subtype.coe_le_coe, subtype.coe_le_coe], + exact ⟨hx, x.2⟩ }, + simp only [subtype.coe_eta, subtype.coe_mk], + end }, end instance nat.locally_finite_order : locally_finite_order ℕ := { finset_Icc := λ a b, (list.range' a (b + 1 - a)).to_finset, coe_finset_Icc := λ a b, begin ext, - rw [finset.mem_coe, list.mem_to_finset, list.mem_range', set.mem_Icc], + rw [mem_coe, list.mem_to_finset, list.mem_range', set.mem_Icc], cases le_or_lt a b, { rw [nat.add_sub_cancel' (nat.lt_succ_of_le h).le, nat.lt_succ_iff] }, rw [nat.sub_eq_zero_iff_le.2 h, add_zero], @@ -279,48 +269,19 @@ instance nat.locally_finite_order : locally_finite_order ℕ := end } instance int.locally_finite_order : locally_finite_order ℤ := -{ finset_Icc := λ a b, (finset.Iio (b + 1 - a).to_nat).map ⟨λ n, n + a, λ _, by simp only +{ finset_Icc := λ a b, (Iio (b + 1 - a).to_nat).map ⟨λ n, n + a, λ _, by simp only [imp_self, forall_const, add_left_inj, int.coe_nat_inj']⟩, coe_finset_Icc := λ a b, begin ext, - rw [finset.mem_coe, finset.mem_map, set.mem_Icc, function.embedding.coe_fn_mk], + rw [mem_coe, mem_map, set.mem_Icc, function.embedding.coe_fn_mk], split, - { - rintro ⟨n, hn, hx⟩, - rw finset.mem_Iio_iff at hn, - rw ←hx, - sorry - }, + { rintro ⟨n, hn, hx⟩, + rw [←hx, le_add_iff_nonneg_left], + rw [mem_Iio_iff, int.lt_to_nat, lt_sub_iff_add_lt, int.lt_add_one_iff] at hn, + exact ⟨int.coe_nat_nonneg _, hn⟩ }, rintro h, - refine ⟨(x - a).to_nat, _, _⟩, - { rw [finset.mem_Iio_iff, ←int.coe_nat_lt, int.to_nat_sub_of_le _ _ h.1], - exact int.to_nat_lt_to_nat (sub_le_sub_right h.2 _) }, - rw sub_eq_add_neg, - rw int.to_nat_sub_of_le _ _ h.1, - simp, - cases le_or_lt a b, - { - rw int.to_nat_sub_of_le _ _ h, - }, - cases le_or_lt a b, - { rw [int.to_nat_of_nonneg (sub_nonneg.2 (h.trans (lt_add_one _).le)), ←add_sub_assoc, - add_sub_cancel', int.lt_add_one_iff] }, - rw [int.to_nat_of_nonpos (sub_nonpos.2 h), int.coe_nat_zero, add_zero], - exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)), - end } - -@[priority 900] -noncomputable instance fintype.locally_finite_order [preorder α] [fintype α] : - locally_finite_order α := -{ finset_Icc := λ a b, (finite.of_fintype (Icc a b)).to_finset, - coe_finset_Icc := λ a b, finite.coe_to_finset _ } - -noncomputable def order_embedding.locally_finite_order {β : Type*} [partial_order β] - [locally_finite_order β] [decidable_eq β] [partial_order α] [decidable_eq α] (f : α ↪o β) : - locally_finite_order α := -{ finset_Icc := λ a b, (finset.Icc (f a) (f b)).preimage f (f.to_embedding.injective.inj_on _), - coe_finset_Icc := λ a b, begin - ext, - simp only [finset.coe_Icc, finset.coe_preimage, iff_self, mem_preimage, mem_Icc, - order_embedding.le_iff_le], + refine ⟨(x - a).to_nat, _, by rw [int.to_nat_sub_of_le h.1, int.sub_add_cancel]⟩, + rw mem_Iio_iff, + have hb := int.lt_add_one_of_le h.2, + exact (int.to_nat_lt_to_nat (sub_pos.2 (h.1.trans_lt hb))).2 (sub_lt_sub_right hb _), end } From c939dda403913ef21098ade5f6302bdaa4d991d4 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Wed, 23 Jun 2021 17:19:11 +0100 Subject: [PATCH 05/41] some of the API --- src/algebra/big_operators/intervals.lean | 6 +- src/analysis/p_series.lean | 6 +- src/data/finset/intervals.lean | 332 +++++++++++++++---- src/data/fintype/intervals.lean | 2 +- src/data/nat/multiplicity.lean | 8 +- src/data/polynomial/div.lean | 2 +- src/data/polynomial/iterated_deriv.lean | 2 +- src/number_theory/divisors.lean | 6 +- src/number_theory/primorial.lean | 7 +- src/number_theory/quadratic_reciprocity.lean | 20 +- src/order/locally_finite.lean | 107 ------ src/topology/metric_space/basic.lean | 2 +- src/topology/metric_space/emetric_space.lean | 2 +- 13 files changed, 294 insertions(+), 208 deletions(-) diff --git a/src/algebra/big_operators/intervals.lean b/src/algebra/big_operators/intervals.lean index 551dcfa798557..a940880bf2397 100644 --- a/src/algebra/big_operators/intervals.lean +++ b/src/algebra/big_operators/intervals.lean @@ -25,7 +25,7 @@ variables {α : Type u} {β : Type v} {γ : Type w} {s₂ s₁ s : finset α} {a lemma sum_Ico_add {δ : Type*} [add_comm_monoid δ] (f : ℕ → δ) (m n k : ℕ) : (∑ l in Ico m n, f (k + l)) = (∑ l in Ico (m + k) (n + k), f l) := -Ico.image_add m n k ▸ eq.symm $ sum_image $ λ x hx y hy h, nat.add_left_cancel h +image_add_const_Ico m n k ▸ eq.symm $ sum_image $ λ x hx y hy h, nat.add_left_cancel h @[to_additive] lemma prod_Ico_add (f : ℕ → β) (m n k : ℕ) : @@ -34,7 +34,7 @@ lemma prod_Ico_add (f : ℕ → β) (m n k : ℕ) : lemma sum_Ico_succ_top {δ : Type*} [add_comm_monoid δ] {a b : ℕ} (hab : a ≤ b) (f : ℕ → δ) : (∑ k in Ico a (b + 1), f k) = (∑ k in Ico a b, f k) + f b := -by rw [Ico.succ_top hab, sum_insert Ico.not_mem_top, add_comm] +by rw [Ico.succ_top hab, sum_insert right_not_mem_Ico, add_comm] @[to_additive] lemma prod_Ico_succ_top {a b : ℕ} (hab : a ≤ b) (f : ℕ → β) : @@ -54,7 +54,7 @@ lemma prod_eq_prod_Ico_succ_bot {a b : ℕ} (hab : a < b) (f : ℕ → β) : @[to_additive] lemma prod_Ico_consecutive (f : ℕ → β) {m n k : ℕ} (hmn : m ≤ n) (hnk : n ≤ k) : (∏ i in Ico m n, f i) * (∏ i in Ico n k, f i) = (∏ i in Ico m k, f i) := -Ico.union_consecutive hmn hnk ▸ eq.symm $ prod_union $ Ico.disjoint_consecutive m n k +Ico_union_Ico_consecutive hmn hnk ▸ eq.symm $ prod_union $ Ico.disjoint_consecutive m n k @[to_additive] lemma prod_range_mul_prod_Ico (f : ℕ → β) {m n : ℕ} (h : m ≤ n) : diff --git a/src/analysis/p_series.lean b/src/analysis/p_series.lean index 8ab5787743743..4d3e406613349 100644 --- a/src/analysis/p_series.lean +++ b/src/analysis/p_series.lean @@ -50,7 +50,7 @@ begin exact add_le_add ihn this, exacts [n.one_le_two_pow, nat.pow_le_pow_of_le_right zero_lt_two n.le_succ] }, have : ∀ k ∈ Ico (2 ^ n) (2 ^ (n + 1)), f k ≤ f (2 ^ n) := - λ k hk, hf (pow_pos zero_lt_two _) (Ico.mem.mp hk).1, + λ k hk, hf (pow_pos zero_lt_two _) (mem_Ico.mp hk).1, convert sum_le_sum this, simp [pow_succ, two_mul] end @@ -72,8 +72,8 @@ begin exacts [add_le_add_right n.one_le_two_pow _, add_le_add_right (nat.pow_le_pow_of_le_right zero_lt_two n.le_succ) _] }, have : ∀ k ∈ Ico (2 ^ n + 1) (2 ^ (n + 1) + 1), f (2 ^ (n + 1)) ≤ f k := - λ k hk, hf (n.one_le_two_pow.trans_lt $ (nat.lt_succ_of_le le_rfl).trans_le (Ico.mem.mp hk).1) - (nat.le_of_lt_succ $ (Ico.mem.mp hk).2), + λ k hk, hf (n.one_le_two_pow.trans_lt $ (nat.lt_succ_of_le le_rfl).trans_le (mem_Ico.mp hk).1) + (nat.le_of_lt_succ $ (mem_Ico.mp hk).2), convert sum_le_sum this, simp [pow_succ, two_mul] end diff --git a/src/data/finset/intervals.lean b/src/data/finset/intervals.lean index 9c2ed00c5f1bc..48ec683460ba3 100644 --- a/src/data/finset/intervals.lean +++ b/src/data/finset/intervals.lean @@ -3,118 +3,312 @@ Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ +import algebra.ordered_monoid import data.finset.basic import data.multiset.intervals +import order.locally_finite /-! -# Intervals in ℕ as finsets +# Intervals as finsets -For now this only covers `Ico n m`, the "closed-open" interval containing `[n, ..., m-1]`. +For now this only covers `Ico a b`, the "closed-open" interval containing `[n, ..., m-1]`. -/ -namespace finset open multiset nat -/-! ### intervals -/ -/- Ico (a closed open interval) -/ -variables {n m l : ℕ} +/-! ### Intervals as finsets -/ + +namespace finset +section preorder +variables {α : Type*} [decidable_eq α] [preorder α] [locally_finite_order α] + +/-- The finset of elements `x` such that `a ≤ x ∧ x ≤ b`. Basically `set.Icc a b` as a finset. -/ +def Icc (a b : α) : finset α := +locally_finite_order.finset_Icc a b + +/- `Icc` is treated separately from `Ico`, `Ioc`, `Ioo` because defining it doesn't require +`decidable_rel (≤)`. Useful in simple constructions, like `subtype.locally_finite_order`. -/ +@[simp, norm_cast] lemma coe_Icc (a b : α) : (Icc a b : set α) = set.Icc a b := +locally_finite_order.coe_finset_Icc a b + +@[simp] lemma mem_Icc {a b x : α} : x ∈ Icc a b ↔ a ≤ x ∧ x ≤ b := +by rw [←set.mem_Icc, ←coe_Icc, mem_coe] + +variable [decidable_rel ((≤) : α → α → Prop)] + +/-- The finset of elements `x` such that `a ≤ x ∧ x < b`. Basically `set.Ico a b` as a finset. -/ +def Ico (a b : α) : finset α := +(Icc a b).filter (λ x, ¬b ≤ x) + +/-- The finset of elements `x` such that `a < x ∧ x ≤ b`. Basically `set.Ioc a b` as a finset. -/ +def Ioc (a b : α) : finset α := +(Icc a b).filter (λ x, ¬x ≤ a) -/-- `Ico n m` is the set of natural numbers `n ≤ k < m`. -/ -def Ico (n m : ℕ) : finset ℕ := ⟨_, Ico.nodup n m⟩ +/-- The finset of elements `x` such that `a < x ∧ x < b`. Basically `set.Ioo a b` as a finset. -/ +def Ioo (a b : α) : finset α := +(Ico a b).filter (λ x, ¬x ≤ a) -namespace Ico +@[simp, norm_cast] lemma coe_Ico (a b : α) : (Ico a b : set α) = set.Ico a b := +by { ext x, rw [mem_coe, Ico, mem_filter, mem_Icc, and_assoc, ←lt_iff_le_not_le, set.mem_Ico] } -@[simp] theorem val (n m : ℕ) : (Ico n m).1 = multiset.Ico n m := rfl +@[simp] lemma mem_Ico {a b x : α} : x ∈ Ico a b ↔ a ≤ x ∧ x < b := +by rw [←set.mem_Ico, ←coe_Ico, mem_coe] -@[simp] theorem to_finset (n m : ℕ) : (multiset.Ico n m).to_finset = Ico n m := -(multiset.to_finset_eq _).symm +@[simp, norm_cast] lemma coe_Ioc (a b : α) : (Ioc a b : set α) = set.Ioc a b := +by { ext x, rw [mem_coe, Ioc, mem_filter, mem_Icc, and.right_comm, ←lt_iff_le_not_le, set.mem_Ioc] } -theorem image_add (n m k : ℕ) : (Ico n m).image ((+) k) = Ico (n + k) (m + k) := -by simp [image, multiset.Ico.map_add] +@[simp] lemma mem_Ioc {a b x : α} : x ∈ Ioc a b ↔ a < x ∧ x ≤ b := +by rw [←set.mem_Ioc, ←coe_Ioc, mem_coe] -theorem image_sub (n m k : ℕ) (h : k ≤ n) : (Ico n m).image (λ x, x - k) = Ico (n - k) (m - k) := +@[simp, norm_cast] lemma coe_Ioo (a b : α) : (Ioo a b : set α) = set.Ioo a b := +by { ext x, rw [mem_coe, Ioo, mem_filter, mem_Ico, and.right_comm, ←lt_iff_le_not_le, set.mem_Ioo] } + +@[simp] lemma mem_Ioo {a b x : α} : x ∈ Ioo a b ↔ a < x ∧ x < b := +by rw [←set.mem_Ioo, ←coe_Ioo, mem_coe] + +theorem Ico_subset_Ico {a₁ b₁ a₂ b₂ : α} (ha : a₂ ≤ a₁) (hb : b₁ ≤ b₂) : + Ico a₁ b₁ ⊆ Ico a₂ b₂ := begin - dsimp [image], - rw [multiset.Ico.map_sub _ _ _ h, ←multiset.to_finset_eq], - refl, + rintro x hx, + rw mem_Ico at ⊢ hx, + exact ⟨ha.trans hx.1, hx.2.trans_le hb⟩, end -theorem zero_bot (n : ℕ) : Ico 0 n = range n := -eq_of_veq $ multiset.Ico.zero_bot _ +end preorder -@[simp] theorem card (n m : ℕ) : (Ico n m).card = m - n := -multiset.Ico.card _ _ +section partial_order +variables {α : Type*} [decidable_eq α] [partial_order α] [decidable_rel ((≤) : α → α → Prop)] + [locally_finite_order α] + +end partial_order + +section order_top +variables {α : Type*} [decidable_eq α] [order_top α] [locally_finite_order α] + +/-- The finset of elements `x` such that `a ≤ x`. Basically `set.Ici a` as a finset. -/ +def Ici (a : α) : finset α := +Icc a ⊤ + +@[simp, norm_cast] lemma coe_Ici (a : α) : (Ici a : set α) = set.Ici a := +by rw [Ici, coe_Icc, set.Icc_top] + +@[simp] lemma mem_Ici {a x : α} : x ∈ Ici a ↔ a ≤ x := +by rw [←set.mem_Ici, ←coe_Ici, mem_coe] + +variable [decidable_rel ((≤) : α → α → Prop)] + +/-- The finset of elements `x` such that `a < x`. Basically `set.Ioi a` as a finset. -/ +def Ioi (a : α) : finset α := +Ioc a ⊤ + +@[simp, norm_cast] lemma coe_Ioi (a : α) : (Ioi a : set α) = set.Ioi a := +by rw [Ioi, coe_Ioc, set.Ioc_top] + +@[simp] lemma mem_Ioi {a x : α} : x ∈ Ioi a ↔ a < x := +by rw [←set.mem_Ioi, ←coe_Ioi, mem_coe] + +end order_top + +section order_bot +variables {α : Type*} [decidable_eq α] [order_bot α] [locally_finite_order α] + +/-- The finset of elements `x` such that `x ≤ b`. Basically `set.Iic b` as a finset. -/ +def Iic (b : α) : finset α := +Icc ⊥ b + +@[simp, norm_cast] lemma coe_Iic (b : α) : (Iic b : set α) = set.Iic b := +by rw [Iic, coe_Icc, set.Icc_bot] + +@[simp] lemma mem_Iic {b x : α} : x ∈ Iic b ↔ x ≤ b := +by rw [←set.mem_Iic, ←coe_Iic, mem_coe] + +variable [decidable_rel ((≤) : α → α → Prop)] + +/-- The finset of elements `x` such that `x < b`. Basically `set.Iio b` as a finset. -/ +def Iio (b : α) : finset α := +Ico ⊥ b + +@[simp, norm_cast] lemma coe_Iio (b : α) : (Iio b : set α) = set.Iio b := +by rw [Iio, coe_Ico, set.Ico_bot] + +@[simp] lemma mem_Iio {b x : α} : x ∈ Iio b ↔ x < b := +by rw [←set.mem_Iio, ←coe_Iio, mem_coe] + +end order_bot + +section ordered_cancel_add_comm_monoid +variables {α : Type*} [decidable_eq α] [ordered_cancel_add_comm_monoid α] [has_exists_add_of_le α] + [locally_finite_order α] -@[simp] theorem mem {n m l : ℕ} : l ∈ Ico n m ↔ n ≤ l ∧ l < m := -multiset.Ico.mem +lemma image_add_const_Icc (a b c : α) : (Icc a b).image ((+) c) = Icc (a + c) (b + c) := +begin + ext x, + rw [mem_image, mem_Icc], + split, + { rintro ⟨y, hy, hx⟩, + rw mem_Icc at hy, + rw [←hx, add_comm c], + exact ⟨add_le_add_right hy.1 c, add_le_add_right hy.2 c⟩ }, + intro hx, + obtain ⟨y, hy⟩ := exists_add_of_le hx.1, + rw [hy, add_right_comm] at hx, + rw [eq_comm, add_right_comm, add_comm] at hy, + exact ⟨a + y, mem_Icc.2 ⟨le_of_add_le_add_right hx.1, le_of_add_le_add_right hx.2⟩, hy⟩, +end + +variable [decidable_rel ((≤) : α → α → Prop)] + +lemma image_add_const_Ico (a b c : α) : (Ico a b).image ((+) c) = Ico (a + c) (b + c) := +begin + ext x, + rw [mem_image, mem_Ico], + split, + { rintro ⟨y, hy, hx⟩, + rw mem_Ico at hy, + rw [←hx, add_comm c], + exact ⟨add_le_add_right hy.1 c, add_lt_add_right hy.2 c⟩ }, + intro hx, + obtain ⟨y, hy⟩ := exists_add_of_le hx.1, + rw [hy, add_right_comm] at hx, + rw [eq_comm, add_right_comm, add_comm] at hy, + exact ⟨a + y, mem_Ico.2 ⟨le_of_add_le_add_right hx.1, lt_of_add_lt_add_right hx.2⟩, hy⟩, +end + +lemma image_add_const_Ioc (a b c : α) : (Ioc a b).image ((+) c) = Ioc (a + c) (b + c) := +begin + ext x, + rw [mem_image, mem_Ioc], + split, + { rintro ⟨y, hy, hx⟩, + rw mem_Ioc at hy, + rw [←hx, add_comm c], + exact ⟨add_lt_add_right hy.1 c, add_le_add_right hy.2 c⟩ }, + intro hx, + obtain ⟨y, hy⟩ := exists_add_of_le hx.1.le, + rw [hy, add_right_comm] at hx, + rw [eq_comm, add_right_comm, add_comm] at hy, + exact ⟨a + y, mem_Ioc.2 ⟨lt_of_add_lt_add_right hx.1, le_of_add_le_add_right hx.2⟩, hy⟩, +end -@[simp, norm_cast] lemma coe_eq_Ico {n m : ℕ} : ↑(Ico n m) = set.Ico n m := +lemma image_add_const_Ioo (a b c : α) : (Ioo a b).image ((+) c) = Ioo (a + c) (b + c) := begin - ext, - rw [set.mem_Ico, mem_coe, finset.Ico.mem], + ext x, + rw [mem_image, mem_Ioo], + split, + { rintro ⟨y, hy, hx⟩, + rw mem_Ioo at hy, + rw [←hx, add_comm c], + exact ⟨add_lt_add_right hy.1 c, add_lt_add_right hy.2 c⟩ }, + intro hx, + obtain ⟨y, hy⟩ := exists_add_of_le hx.1.le, + rw [hy, add_right_comm] at hx, + rw [eq_comm, add_right_comm, add_comm] at hy, + exact ⟨a + y, mem_Ioo.2 ⟨lt_of_add_lt_add_right hx.1, lt_of_add_lt_add_right hx.2⟩, hy⟩, end -theorem eq_empty_of_le {n m : ℕ} (h : m ≤ n) : Ico n m = ∅ := +lemma image_sub (a b c : α) (h : a ≤ c) : (Ico a b).image (λ x, x - c) = Ico (a - c) (b - c) := +begin + dsimp [image], + rw [multiset.Ico.map_sub _ _ _ h, ←multiset.to_finset_eq], + refl, +end + +lemma Ico_eq_range' (a b : ℕ) : Ico a b = (list.range' a (b - a)).to_finset := +begin + ext x, + rw [mem_Ico, list.mem_to_finset, list.mem_range'], + cases le_total a b, + { rw nat.add_sub_cancel' h }, + rw [nat.sub_eq_zero_iff_le.2 h, add_zero], + exact iff_of_false (λ hx, (h.trans hx.1).not_lt hx.2) (λ hx, hx.1.not_lt hx.2), +end + +-- TODO: Once we have `has_lt_iff_add_one_le`, we can generalise that +/-- Currently only for ℕ -/ +@[simp] lemma card_Ico (a b : ℕ) : (Ico a b).card = a - b := +multiset.Ico.card _ _ + +theorem eq_empty_of_le {n m : ℕ} (h : m ≤ n) : Ico a b = ∅ := eq_of_veq $ multiset.Ico.eq_zero_of_le h @[simp] theorem self_eq_empty (n : ℕ) : Ico n n = ∅ := eq_empty_of_le $ le_refl n -@[simp] theorem eq_empty_iff {n m : ℕ} : Ico n m = ∅ ↔ m ≤ n := -iff.trans val_eq_zero.symm multiset.Ico.eq_zero_iff +@[simp] theorem Ico_nonempty_iff {a b : α} : (Ico a b).nonempty ↔ a < b := +begin + rw ←coe_nonempty, + rw set.Ico_eq, + split, + { + rintro ⟨x, hx⟩, + rw mem_Ico at hx, + } +end -theorem subset_iff {m₁ n₁ m₂ n₂ : ℕ} (hmn : m₁ < n₁) : - Ico m₁ n₁ ⊆ Ico m₂ n₂ ↔ (m₂ ≤ m₁ ∧ n₁ ≤ n₂) := +class has_lt_iff_add_one_le (α : Type*) [preorder α] [has_add α] [has_one α] : + Prop := +(lt_iff_add_one_le : ∀ a b : α, a < b ↔ a + 1 ≤ b) + +@[simp] theorem Ico_eq_empty_iff {a b : α} : Ico a b = ∅ ↔ ¬a < b := +by { rw [←not_nonempty_iff_eq_empty, not_iff_not], exact Ico_nonempty_iff } + +theorem Ico_subset_Ico_iff {a₁ b₁ a₂ b₂ : α} (h : a₁ < b₁) : + Ico a₁ b₁ ⊆ Ico a₂ b₂ ↔ (a₂ ≤ a₁ ∧ b₁ ≤ b₂) := begin - simp only [subset_iff, mem], - refine ⟨λ h, ⟨_, _⟩, _⟩, - { exact (h ⟨le_refl _, hmn⟩).1 }, + simp_rw [subset_iff, mem_Ico], + refine ⟨λ hx, ⟨(hx ⟨le_rfl, h⟩).1, _⟩, _⟩, + { exact (hx ⟨le_rfl, h⟩).1 }, { refine le_of_pred_lt (@h (pred n₁) ⟨le_pred_of_lt hmn, pred_lt _⟩).2, exact ne_of_gt (lt_of_le_of_lt (nat.zero_le m₁) hmn) }, { rintros ⟨hm, hn⟩ k ⟨hmk, hkn⟩, exact ⟨le_trans hm hmk, lt_of_lt_of_le hkn hn⟩ } end -protected theorem subset {m₁ n₁ m₂ n₂ : ℕ} (hmm : m₂ ≤ m₁) (hnn : n₁ ≤ n₂) : - Ico m₁ n₁ ⊆ Ico m₂ n₂ := +end ordered_cancel_add_comm_monoid + +section ordered_semiring +variables {α : Type*} [decidable_eq α] [ordered_semiring α] [locally_finite_order α] + +end ordered_semiring + +section linear_order +variables {α : Type*} [decidable_eq α] [linear_order α] [locally_finite_order α] + +lemma Ico_union_Ico_eq_Ico {a b c : α} (hab : a ≤ b) (hbc : b ≤ c) : + Ico a b ∪ Ico b c = Ico a c := begin - simp only [finset.subset_iff, Ico.mem], - assume x hx, - exact ⟨le_trans hmm hx.1, lt_of_lt_of_le hx.2 hnn⟩ + ext x, + rw [mem_union, ←mem_coe, ←mem_coe, ←mem_coe, coe_Ico, coe_Ico, coe_Ico, ←set.mem_union, + set.Ico_union_Ico_eq_Ico hab hbc], end -lemma union_consecutive {n m l : ℕ} (hnm : n ≤ m) (hml : m ≤ l) : - Ico n m ∪ Ico m l = Ico n l := -by rw [← to_finset, ← to_finset, ← multiset.to_finset_add, - multiset.Ico.add_consecutive hnm hml, to_finset] - -lemma union' {n m l k : ℕ} (hlm : l ≤ m) (hnk : n ≤ k) : - Ico n m ∪ Ico l k = Ico (min n l) (max m k) := +lemma Ico_union_Ico' {a b c d : α} (hcb : c ≤ b) (had : a ≤ d) : + Ico a b ∪ Ico c d = Ico (min n l) (max m k) := by simp [←coe_inj, set.Ico_union_Ico' hlm hnk] -lemma union {n m l k : ℕ} (h₁ : min n m ≤ max l k) (h₂ : min l k ≤ max n m) : - Ico n m ∪ Ico l k = Ico (min n l) (max m k) := +lemma union {a b c d : α} (h₁ : min a b ≤ max c d) (h₂ : min c d ≤ max a b) : + Ico a b ∪ Ico l k = Ico (min n l) (max m k) := by simp [←coe_inj, set.Ico_union_Ico h₁ h₂] -@[simp] lemma inter_consecutive (n m l : ℕ) : Ico n m ∩ Ico m l = ∅ := +@[simp] lemma inter_consecutive (n m l : ℕ) : Ico a b ∩ Ico m l = ∅ := begin rw [← to_finset, ← to_finset, ← multiset.to_finset_inter, multiset.Ico.inter_consecutive], simp, end -lemma inter {n m l k : ℕ} : Ico n m ∩ Ico l k = Ico (max n l) (min m k) := +lemma inter {n m l k : ℕ} : Ico a b ∩ Ico l k = Ico (max n l) (min m k) := by simp [←coe_inj, ←inf_eq_min, ←sup_eq_max, set.Ico_inter_Ico] -lemma disjoint_consecutive (n m l : ℕ) : disjoint (Ico n m) (Ico m l) := +lemma disjoint_consecutive (a b l : ℕ) : disjoint (Ico a b) (Ico m l) := le_of_eq $ inter_consecutive n m l @[simp] theorem succ_singleton (n : ℕ) : Ico n (n+1) = {n} := eq_of_veq $ multiset.Ico.succ_singleton -theorem succ_top {n m : ℕ} (h : n ≤ m) : Ico n (m + 1) = insert m (Ico n m) := +theorem succ_top {n m : ℕ} (h : n ≤ m) : Ico n (m + 1) = insert m (Ico a b) := by rw [← to_finset, multiset.Ico.succ_top h, multiset.to_finset_cons, to_finset] -theorem succ_top' {n m : ℕ} (h : n < m) : Ico n m = insert (m - 1) (Ico n (m - 1)) := +theorem succ_top' {n m : ℕ} (h : n < m) : Ico a b = insert (m - 1) (Ico n (m - 1)) := begin have w : m = m - 1 + 1 := (nat.sub_add_cancel (nat.one_le_of_lt h)).symm, conv { to_lhs, rw w }, @@ -122,46 +316,46 @@ begin exact nat.le_pred_of_lt h end -theorem insert_succ_bot {n m : ℕ} (h : n < m) : insert n (Ico (n + 1) m) = Ico n m := +theorem insert_succ_bot {n m : ℕ} (h : n < m) : insert n (Ico (n + 1) m) = Ico a b := by rw [eq_comm, ← to_finset, multiset.Ico.eq_cons h, multiset.to_finset_cons, to_finset] @[simp] theorem pred_singleton {m : ℕ} (h : 0 < m) : Ico (m - 1) m = {m - 1} := eq_of_veq $ multiset.Ico.pred_singleton h -@[simp] theorem not_mem_top {n m : ℕ} : m ∉ Ico n m := -multiset.Ico.not_mem_top +@[simp] theorem right_not_mem_Ico {a b : α} : b ∉ Ico a b := +by { rw [mem_Ico, not_and], exact λ _, lt_irrefl _ } -lemma filter_lt_of_top_le {n m l : ℕ} (hml : m ≤ l) : (Ico n m).filter (λ x, x < l) = Ico n m := +lemma filter_lt_of_top_le {n m l : ℕ} (hml : m ≤ l) : (Ico a b).filter (λ x, x < l) = Ico a b := eq_of_veq $ multiset.Ico.filter_lt_of_top_le hml -lemma filter_lt_of_le_bot {n m l : ℕ} (hln : l ≤ n) : (Ico n m).filter (λ x, x < l) = ∅ := +lemma filter_lt_of_le_bot {n m l : ℕ} (hln : l ≤ n) : (Ico a b).filter (λ x, x < l) = ∅ := eq_of_veq $ multiset.Ico.filter_lt_of_le_bot hln -lemma filter_Ico_bot {n m : ℕ} (hnm : n < m) : (Ico n m).filter (λ x, x ≤ n) = {n} := +lemma filter_Ico_bot {n m : ℕ} (hnm : n < m) : (Ico a b).filter (λ x, x ≤ n) = {n} := eq_of_veq $ multiset.Ico.filter_le_of_bot hnm -lemma filter_lt_of_ge {n m l : ℕ} (hlm : l ≤ m) : (Ico n m).filter (λ x, x < l) = Ico n l := +lemma filter_lt_of_ge {n m l : ℕ} (hlm : l ≤ m) : (Ico a b).filter (λ x, x < l) = Ico n l := eq_of_veq $ multiset.Ico.filter_lt_of_ge hlm -@[simp] lemma filter_lt (n m l : ℕ) : (Ico n m).filter (λ x, x < l) = Ico n (min m l) := +@[simp] lemma filter_lt (n m l : ℕ) : (Ico a b).filter (λ x, x < l) = Ico n (min m l) := eq_of_veq $ multiset.Ico.filter_lt n m l -lemma filter_le_of_le_bot {n m l : ℕ} (hln : l ≤ n) : (Ico n m).filter (λ x, l ≤ x) = Ico n m := +lemma filter_le_of_le_bot {n m l : ℕ} (hln : l ≤ n) : (Ico a b).filter (λ x, l ≤ x) = Ico a b := eq_of_veq $ multiset.Ico.filter_le_of_le_bot hln -lemma filter_le_of_top_le {n m l : ℕ} (hml : m ≤ l) : (Ico n m).filter (λ x, l ≤ x) = ∅ := +lemma filter_le_of_top_le {n m l : ℕ} (hml : m ≤ l) : (Ico a b).filter (λ x, l ≤ x) = ∅ := eq_of_veq $ multiset.Ico.filter_le_of_top_le hml -lemma filter_le_of_le {n m l : ℕ} (hnl : n ≤ l) : (Ico n m).filter (λ x, l ≤ x) = Ico l m := +lemma filter_le_of_le {n m l : ℕ} (hnl : n ≤ l) : (Ico a b).filter (λ x, l ≤ x) = Ico l m := eq_of_veq $ multiset.Ico.filter_le_of_le hnl -@[simp] lemma filter_le (n m l : ℕ) : (Ico n m).filter (λ x, l ≤ x) = Ico (max n l) m := +@[simp] lemma filter_le (n m l : ℕ) : (Ico a b).filter (λ x, l ≤ x) = Ico (max n l) m := eq_of_veq $ multiset.Ico.filter_le n m l -@[simp] lemma diff_left (l n m : ℕ) : (Ico n m) \ (Ico n l) = Ico (max n l) m := +@[simp] lemma diff_left (l n m : ℕ) : (Ico a b) \ (Ico n l) = Ico (max n l) m := by ext k; by_cases n ≤ k; simp [h, and_comm] -@[simp] lemma diff_right (l n m : ℕ) : (Ico n m) \ (Ico l m) = Ico n (min m l) := +@[simp] lemma diff_right (l n m : ℕ) : (Ico a b) \ (Ico l m) = Ico n (min m l) := have ∀k, (k < m ∧ (l ≤ k → m ≤ k)) ↔ (k < m ∧ k < l) := assume k, and_congr_right $ assume hk, by rw [← not_imp_not]; simp [hk], by ext k; by_cases n ≤ k; simp [h, this] @@ -186,8 +380,6 @@ begin { exact nat.sub_sub_self hj } } end -end Ico - lemma range_eq_Ico (n : ℕ) : finset.range n = finset.Ico 0 n := by { ext i, simp } diff --git a/src/data/fintype/intervals.lean b/src/data/fintype/intervals.lean index aa6c98fa36a60..6f2194ab65b94 100644 --- a/src/data/fintype/intervals.lean +++ b/src/data/fintype/intervals.lean @@ -11,7 +11,7 @@ import data.pnat.intervals # fintype instances for intervals We provide `fintype` instances for `Ico l u`, for `l u : ℕ`, and for `l u : ℤ`. --/ +-/a < b ↔ a + 1 ≤ b namespace set diff --git a/src/data/nat/multiplicity.lean b/src/data/nat/multiplicity.lean index bbbc4bfc0b283..6206c90fc13c7 100644 --- a/src/data/nat/multiplicity.lean +++ b/src/data/nat/multiplicity.lean @@ -43,7 +43,7 @@ calc multiplicity m n = ↑(Ico 1 $ ((multiplicity m n).get (finite_nat_iff.2 else mt (le_of_dvd hn0) (not_le_of_lt $ pow_succ_log_gt_self m n (hm1.symm.le_iff_lt.mp (zero_lt_iff.mpr hm0.intro)) hn0), ⟨λ hi, begin - simp only [Ico.mem, mem_filter, lt_succ_iff] at *, + simp only [mem_Ico, mem_filter, lt_succ_iff] at *, exact ⟨⟨hi.1, lt_of_le_of_lt hi.2 $ lt_of_lt_of_le (by rw [← enat.coe_lt_coe, enat.coe_get, multiplicity_lt_iff_neg_dvd]; exact hmn) @@ -52,7 +52,7 @@ calc multiplicity m n = ↑(Ico 1 $ ((multiplicity m n).get (finite_nat_iff.2 rw [← @enat.coe_le_coe i, enat.coe_get] at hi; exact hi.2⟩ end, begin - simp only [Ico.mem, mem_filter, lt_succ_iff, and_imp, true_and] { contextual := tt }, + simp only [mem_Ico, mem_filter, lt_succ_iff, and_imp, true_and] { contextual := tt }, assume h1i hib hmin, rwa [← enat.coe_le_coe, enat.coe_get, ← pow_dvd_iff_le_multiplicity] end⟩ @@ -108,7 +108,7 @@ begin revert hm, have h4 : ∀ m ∈ Ico (p * n + 1) (p * (n + 1)), multiplicity p m = 0, { intros m hm, apply multiplicity_eq_zero_of_not_dvd, - rw [← exists_lt_and_lt_iff_not_dvd _ (pos_iff_ne_zero.mpr hp.ne_zero)], rw [Ico.mem] at hm, + rw [← exists_lt_and_lt_iff_not_dvd _ (pos_iff_ne_zero.mpr hp.ne_zero)], rw [mem_Ico] at hm, exact ⟨n, lt_of_succ_le hm.1, hm.2⟩ }, simp_rw [← prod_Ico_id_eq_factorial, multiplicity.finset.prod hp', ← sum_Ico_consecutive _ h1 h3, add_assoc], intro h, @@ -211,7 +211,7 @@ le_antisymm ← enat.coe_add, enat.coe_le_coe, log_pow _ _ hp.one_lt, ← card_disjoint_union hdisj, filter_union_right], have filter_le_Ico := (Ico 1 n.succ).card_filter_le _, - rwa Ico.card 1 n.succ at filter_le_Ico, + rwa card_Ico 1 n.succ at filter_le_Ico, end) (by rw [← hp.multiplicity_pow_self]; exact multiplicity_le_multiplicity_choose_add hp _ _) diff --git a/src/data/polynomial/div.lean b/src/data/polynomial/div.lean index 39ef3af90ed4f..8f9d4448c5f8c 100644 --- a/src/data/polynomial/div.lean +++ b/src/data/polynomial/div.lean @@ -35,7 +35,7 @@ def div_X (p : polynomial R) : polynomial R := @[simp] lemma coeff_div_X : (div_X p).coeff n = p.coeff (n+1) := begin simp only [div_X, coeff_monomial, true_and, finset_sum_coeff, not_lt, - Ico.mem, zero_le, finset.sum_ite_eq', ite_eq_left_iff], + mem_Ico, zero_le, finset.sum_ite_eq', ite_eq_left_iff], intro h, rw coeff_eq_zero_of_nat_degree_lt (nat.lt_succ_of_le h) end diff --git a/src/data/polynomial/iterated_deriv.lean b/src/data/polynomial/iterated_deriv.lean index 627dbf083b921..eda2ac784751b 100644 --- a/src/data/polynomial/iterated_deriv.lean +++ b/src/data/polynomial/iterated_deriv.lean @@ -138,7 +138,7 @@ begin apply congr_arg2, { refl }, { simp only [prod_singleton], norm_cast }, - { simp only [succ_pos', disjoint_singleton, and_true, lt_add_iff_pos_right, not_le, Ico.mem], + { simp only [succ_pos', disjoint_singleton, and_true, lt_add_iff_pos_right, not_le, mem_Ico], exact lt_add_one (m + 1) } }, { exact congr_arg _ (succ_add m k) } }, end diff --git a/src/number_theory/divisors.lean b/src/number_theory/divisors.lean index 28d237a53e5b1..fd49989f374c4 100644 --- a/src/number_theory/divisors.lean +++ b/src/number_theory/divisors.lean @@ -58,7 +58,7 @@ end @[simp] lemma mem_proper_divisors {m : ℕ} : n ∈ proper_divisors m ↔ n ∣ m ∧ n < m := begin - rw [proper_divisors, finset.mem_filter, finset.Ico.mem, and_comm], + rw [proper_divisors, finset.mem_filter, finset.mem_Ico, and_comm], apply and_congr_right, rw and_iff_right_iff_imp, intros hdvd hlt, @@ -78,7 +78,7 @@ lemma mem_divisors {m : ℕ} : begin cases m, { simp [divisors] }, - simp only [divisors, finset.Ico.mem, ne.def, finset.mem_filter, succ_ne_zero, and_true, + simp only [divisors, finset.mem_Ico, ne.def, finset.mem_filter, succ_ne_zero, and_true, and_iff_right_iff_imp, not_false_iff], intro hdvd, split, @@ -101,7 +101,7 @@ end lemma mem_divisors_antidiagonal {x : ℕ × ℕ} : x ∈ divisors_antidiagonal n ↔ x.fst * x.snd = n ∧ n ≠ 0 := begin - simp only [divisors_antidiagonal, finset.Ico.mem, ne.def, finset.mem_filter, finset.mem_product], + simp only [divisors_antidiagonal, finset.mem_Ico, ne.def, finset.mem_filter, finset.mem_product], rw and_comm, apply and_congr_right, rintro rfl, diff --git a/src/number_theory/primorial.lean b/src/number_theory/primorial.lean index 87e204ad0791d..2a3e08dd573ca 100644 --- a/src/number_theory/primorial.lean +++ b/src/number_theory/primorial.lean @@ -115,7 +115,8 @@ lemma primorial_le_4_pow : ∀ (n : ℕ), n# ≤ 4 ^ n = ∏ i in filter prime (range (2 * m + 2)), i : by simpa [←twice_m] ... = ∏ i in filter prime (finset.Ico (m + 2) (2 * m + 2) ∪ range (m + 2)), i : begin - rw [range_eq_Ico, range_eq_Ico, finset.union_comm, finset.Ico.union_consecutive], + rw [range_eq_Ico, range_eq_Ico, finset.union_comm, + finset.Ico_union_Ico_consecutive], exact bot_le, simp only [add_le_add_iff_right], linarith, @@ -128,7 +129,7 @@ lemma primorial_le_4_pow : ∀ (n : ℕ), n# ≤ 4 ^ n begin apply finset.prod_union, have disj : disjoint (finset.Ico (m + 2) (2 * m + 2)) (range (m + 2)), - { simp only [finset.disjoint_left, and_imp, finset.Ico.mem, not_lt, + { simp only [finset.disjoint_left, and_imp, finset.mem_Ico, not_lt, finset.mem_range], intros _ pr _, exact pr, }, exact finset.disjoint_filter_filter disj, @@ -144,7 +145,7 @@ lemma primorial_le_4_pow : ∀ (n : ℕ), n# ≤ 4 ^ n { intros a, rw finset.mem_filter, intros pr, rcases pr with ⟨ size, is_prime ⟩, - simp only [finset.Ico.mem] at size, + simp only [finset.mem_Ico] at size, rcases size with ⟨ a_big , a_small ⟩, exact dvd_choose_of_middling_prime a is_prime m a_big (nat.lt_succ_iff.mp a_small), }, }, diff --git a/src/number_theory/quadratic_reciprocity.lean b/src/number_theory/quadratic_reciprocity.lean index a9e7d3be9cd2d..39759355db620 100644 --- a/src/number_theory/quadratic_reciprocity.lean +++ b/src/number_theory/quadratic_reciprocity.lean @@ -112,7 +112,7 @@ begin symmetry, refine prod_bij (λ a _, (a : zmod p).val) _ _ _ _, { intros a ha, - rw [Ico.mem, ← nat.succ_sub hp, nat.succ_sub_one], + rw [mem_Ico, ← nat.succ_sub hp, nat.succ_sub_one], split, { apply nat.pos_of_ne_zero, rw ← @val_zero p, assume h, apply units.ne_zero a (val_injective p h) }, @@ -120,7 +120,7 @@ begin { intros a ha, simp only [cast_id, nat_cast_val], }, { intros _ _ _ _ h, rw units.ext_iff, exact val_injective p h }, { intros b hb, - rw [Ico.mem, nat.succ_le_iff, ← succ_sub hp, succ_sub_one, pos_iff_ne_zero] at hb, + rw [mem_Ico, nat.succ_le_iff, ← succ_sub hp, succ_sub_one, pos_iff_ne_zero] at hb, refine ⟨units.mk0 b _, finset.mem_univ _, _⟩, { assume h, apply hb.1, apply_fun val at h, simpa only [val_cast_of_lt hb.right, val_zero] using h }, @@ -157,7 +157,7 @@ begin have hsurj : ∀ (b : ℕ) (hb : b ∈ Ico 1 (p / 2).succ), ∃ x ∈ Ico 1 (p / 2).succ, b = (a * x : zmod p).val_min_abs.nat_abs, { assume b hb, - refine ⟨(b / a : zmod p).val_min_abs.nat_abs, Ico.mem.mpr ⟨_, _⟩, _⟩, + refine ⟨(b / a : zmod p).val_min_abs.nat_abs, mem_Ico.mpr ⟨_, _⟩, _⟩, { apply nat.pos_of_ne_zero, simp only [div_eq_mul_inv, hap, char_p.cast_eq_zero_iff (zmod p) p, hpe hb, not_false_iff, val_min_abs_eq_zero, inv_eq_zero, int.nat_abs_eq_zero, ne.def, mul_eq_zero, or_self] }, @@ -165,9 +165,9 @@ begin { rw nat_cast_nat_abs_val_min_abs, split_ifs, { erw [mul_div_cancel' _ hap, val_min_abs_def_pos, val_cast_of_lt (hep hb), - if_pos (le_of_lt_succ (Ico.mem.1 hb).2), int.nat_abs_of_nat], }, + if_pos (le_of_lt_succ (mem_Ico.1 hb).2), int.nat_abs_of_nat], }, { erw [mul_neg_eq_neg_mul_symm, mul_div_cancel' _ hap, nat_abs_val_min_abs_neg, - val_min_abs_def_pos, val_cast_of_lt (hep hb), if_pos (le_of_lt_succ (Ico.mem.1 hb).2), + val_min_abs_def_pos, val_cast_of_lt (hep hb), if_pos (le_of_lt_succ (mem_Ico.1 hb).2), int.nat_abs_of_nat] } } }, exact multiset.map_eq_map_of_bij_of_nodup _ _ (finset.nodup _) (finset.nodup _) (λ x _, (a * x : zmod p).val_min_abs.nat_abs) hmem (λ _ _, rfl) @@ -182,7 +182,7 @@ private lemma gauss_lemma_aux₁ (p : ℕ) [fact p.prime] [fact (p % 2 = 1)] calc (a ^ (p / 2) * (p / 2)! : zmod p) = (∏ x in Ico 1 (p / 2).succ, a * x) : by rw [prod_mul_distrib, ← prod_nat_cast, ← prod_nat_cast, prod_Ico_id_eq_factorial, - prod_const, Ico.card, succ_sub_one]; simp + prod_const, card_Ico, succ_sub_one]; simp ... = (∏ x in Ico 1 (p / 2).succ, (a * x : zmod p).val) : by simp ... = (∏ x in Ico 1 (p / 2).succ, (if (a * x : zmod p).val ≤ p / 2 then 1 else -1) * @@ -325,23 +325,23 @@ begin (λ x : ℕ × ℕ, x.1 * q ≤ x.2 * p)), { apply disjoint_filter.2 (λ x hx hpq hqp, _), have hxp : x.1 < p, from lt_of_le_of_lt - (show x.1 ≤ p / 2, by simp only [*, lt_succ_iff, Ico.mem, mem_product] at *; tauto) + (show x.1 ≤ p / 2, by simp only [*, lt_succ_iff, mem_Ico, mem_product] at *; tauto) (nat.div_lt_self hp.1.pos dec_trivial), have : (x.1 : zmod p) = 0, { simpa [hq0] using congr_arg (coe : ℕ → zmod p) (le_antisymm hpq hqp) }, apply_fun zmod.val at this, rw [val_cast_of_lt hxp, val_zero] at this, - simpa only [this, nonpos_iff_eq_zero, Ico.mem, one_ne_zero, false_and, mem_product] using hx }, + simpa only [this, nonpos_iff_eq_zero, mem_Ico, one_ne_zero, false_and, mem_product] using hx }, have hunion : ((Ico 1 (p / 2).succ).product (Ico 1 (q / 2).succ)).filter (λ x : ℕ × ℕ, x.2 * p ≤ x.1 * q) ∪ ((Ico 1 (p / 2).succ).product (Ico 1 (q / 2).succ)).filter (λ x : ℕ × ℕ, x.1 * q ≤ x.2 * p) = ((Ico 1 (p / 2).succ).product (Ico 1 (q / 2).succ)), from finset.ext (λ x, by have := le_total (x.2 * p) (x.1 * q); - simp only [mem_union, mem_filter, Ico.mem, mem_product]; tauto), + simp only [mem_union, mem_filter, mem_Ico, mem_product]; tauto), rw [sum_Ico_eq_card_lt, sum_Ico_eq_card_lt, hswap, ← card_disjoint_union hdisj, hunion, card_product], - simp only [Ico.card, nat.sub_zero, succ_sub_succ_eq_sub] + simp only [card_Ico, nat.sub_zero, succ_sub_succ_eq_sub] end variables (p q : ℕ) [fact p.prime] [fact q.prime] diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index e5010c7c34644..6fdad6fa19df6 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -97,113 +97,6 @@ lemma Ioo_finite (a b : α) : (Ioo a b).finite := end set -/-! ### Intervals as finsets -/ - -namespace finset -section preorder -variables {α : Type*} [decidable_eq α] [preorder α] [locally_finite_order α] - -/-- The finset of elements `x` such that `a ≤ x ∧ x ≤ b`. Basically `set.Icc a b` as a finset. -/ -def Icc (a b : α) : finset α := -locally_finite_order.finset_Icc a b - -/-- The finset of elements `x` such that `a ≤ x ∧ x < b`. Basically `set.Ico a b` as a finset. -/ -def Ico (a b : α) : finset α := -(Icc a b).erase b - -/-- The finset of elements `x` such that `a < x ∧ x ≤ b`. Basically `set.Ioc a b` as a finset. -/ -def Ioc (a b : α) : finset α := -(Icc a b).erase a - -/-- The finset of elements `x` such that `a < x ∧ x < b`. Basically `set.Ioo a b` as a finset. -/ -def Ioo (a b : α) : finset α := -(Ioc a b).erase b - -/- `Icc` is treated separately from `Ico`, `Ioc`, `Ioo` because it allows to relax `partial_order` -to `preorder` in simple constructions, eg `subtype.locally_finite_order`. -/ -@[simp] lemma coe_Icc (a b : α) : (Icc a b : set α) = set.Icc a b := -locally_finite_order.coe_finset_Icc a b - -@[simp] lemma mem_Icc_iff {a b x : α} : x ∈ Icc a b ↔ a ≤ x ∧ x ≤ b := -by rw [←set.mem_Icc, ←coe_Icc, mem_coe] - -end preorder - -section partial_order -variables {α : Type*} [decidable_eq α] [partial_order α] [locally_finite_order α] - -@[simp] lemma coe_Ico (a b : α) : (Ico a b : set α) = set.Ico a b := -by rw [Ico, coe_erase, coe_Icc, set.Icc_diff_right] - -@[simp] lemma coe_Ioc (a b : α) : (Ioc a b : set α) = set.Ioc a b := -by rw [Ioc, coe_erase, coe_Icc, set.Icc_diff_left] - -@[simp] lemma coe_Ioo (a b : α) : (Ioo a b : set α) = set.Ioo a b := -by rw [Ioo, coe_erase, coe_Ioc, set.Ioc_diff_right] - -@[simp] lemma mem_Ico_iff {a b x : α} : x ∈ Ico a b ↔ a ≤ x ∧ x < b := -by rw [←set.mem_Ico, ←coe_Ico, mem_coe] - -@[simp] lemma mem_Ioc_iff {a b x : α} : x ∈ Ioc a b ↔ a < x ∧ x ≤ b := -by rw [←set.mem_Ioc, ←coe_Ioc, mem_coe] - -@[simp] lemma mem_Ioo_iff {a b x : α} : x ∈ Ioo a b ↔ a < x ∧ x < b := -by rw [←set.mem_Ioo, ←coe_Ioo, mem_coe] - -end partial_order - -section order_top -variables {α : Type*} [decidable_eq α] [order_top α] [locally_finite_order α] - -/-- The finset of elements `x` such that `a ≤ x`. Basically `set.Ici a` as a finset. -/ -def Ici (a : α) : finset α := -Icc a ⊤ - -/-- The finset of elements `x` such that `a < x`. Basically `set.Ioi a` as a finset. -/ -def Ioi (a : α) : finset α := -Ioc a ⊤ - -@[simp] lemma coe_Ici (a : α) : (Ici a : set α) = set.Ici a := -by rw [Ici, coe_Icc, set.Icc_top] - -@[simp] lemma coe_Ioi (a : α) : (Ioi a : set α) = set.Ioi a := -by rw [Ioi, coe_Ioc, set.Ioc_top] - -@[simp] lemma mem_Ici_iff {a x : α} : x ∈ Ici a ↔ a ≤ x := -by rw [←set.mem_Ici, ←coe_Ici, mem_coe] - -@[simp] lemma mem_Ioi_iff {a x : α} : x ∈ Ioi a ↔ a < x := -by rw [←set.mem_Ioi, ←coe_Ioi, mem_coe] - -end order_top - -section order_bot -variables {α : Type*} [decidable_eq α] [order_bot α] [locally_finite_order α] - -/-- The finset of elements `x` such that `x ≤ b`. Basically `set.Iic b` as a finset. -/ -def Iic (b : α) : finset α := -Icc ⊥ b - -/-- The finset of elements `x` such that `x < b`. Basically `set.Iio b` as a finset. -/ -def Iio (b : α) : finset α := -Ico ⊥ b - -@[simp] lemma coe_Iic (b : α) : (Iic b : set α) = set.Iic b := -by rw [Iic, coe_Icc, set.Icc_bot] - -@[simp] lemma coe_Iio (b : α) : (Iio b : set α) = set.Iio b := -by rw [Iio, coe_Ico, set.Ico_bot] - -@[simp] lemma mem_Iic_iff {b x : α} : x ∈ Iic b ↔ x ≤ b := -by rw [←set.mem_Iic, ←coe_Iic, mem_coe] - -@[simp] lemma mem_Iio_iff {b x : α} : x ∈ Iio b ↔ x < b := -by rw [←set.mem_Iio, ←coe_Iio, mem_coe] - -end order_bot - -end finset - open finset /-! ### Instances -/ diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index b1079d24cfe83..b8c307727159a 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -165,7 +165,7 @@ lemma dist_le_Ico_sum_of_dist_le {f : ℕ → α} {m n} (hmn : m ≤ n) {d : ℕ → ℝ} (hd : ∀ {k}, m ≤ k → k < n → dist (f k) (f (k + 1)) ≤ d k) : dist (f m) (f n) ≤ ∑ i in finset.Ico m n, d i := le_trans (dist_le_Ico_sum_dist f hmn) $ -finset.sum_le_sum $ λ k hk, hd (finset.Ico.mem.1 hk).1 (finset.Ico.mem.1 hk).2 +finset.sum_le_sum $ λ k hk, hd (finset.mem_Ico.1 hk).1 (finset.mem_Ico.1 hk).2 /-- A version of `dist_le_range_sum_dist` with each intermediate distance replaced with an upper estimate. -/ diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index 15f1883ba55c4..de6f4a4fbb0e2 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -153,7 +153,7 @@ lemma edist_le_Ico_sum_of_edist_le {f : ℕ → α} {m n} (hmn : m ≤ n) {d : ℕ → ℝ≥0∞} (hd : ∀ {k}, m ≤ k → k < n → edist (f k) (f (k + 1)) ≤ d k) : edist (f m) (f n) ≤ ∑ i in finset.Ico m n, d i := le_trans (edist_le_Ico_sum_edist f hmn) $ -finset.sum_le_sum $ λ k hk, hd (finset.Ico.mem.1 hk).1 (finset.Ico.mem.1 hk).2 +finset.sum_le_sum $ λ k hk, hd (finset.mem_Ico.1 hk).1 (finset.mem_Ico.1 hk).2 /-- A version of `edist_le_range_sum_edist` with each intermediate distance replaced with an upper estimate. -/ From e9fdb7567ee05776748a438e74c775d8823f86f7 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Tue, 14 Sep 2021 11:48:46 +0200 Subject: [PATCH 06/41] reorder --- src/data/finset/intervals.lean | 155 +------------ src/data/fintype/intervals.lean | 76 ------ src/data/int/intervals.lean | 79 +++++++ src/data/nat/intervals.lean | 51 ++++ src/data/pnat/intervals.lean | 56 ++++- src/order/locally_finite.lean | 397 ++++++++++++++++++++++++++------ 6 files changed, 507 insertions(+), 307 deletions(-) delete mode 100644 src/data/fintype/intervals.lean create mode 100644 src/data/int/intervals.lean create mode 100644 src/data/nat/intervals.lean diff --git a/src/data/finset/intervals.lean b/src/data/finset/intervals.lean index 48ec683460ba3..c5fef0ed48143 100644 --- a/src/data/finset/intervals.lean +++ b/src/data/finset/intervals.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import algebra.ordered_monoid +import algebra.ordered_sub import data.finset.basic import data.multiset.intervals import order.locally_finite @@ -11,134 +11,14 @@ import order.locally_finite /-! # Intervals as finsets -For now this only covers `Ico a b`, the "closed-open" interval containing `[n, ..., m-1]`. +For now this only covers `Ico a b`, the "closed-open" interval containing `[a, ..., b - 1]`. -/ -open multiset nat - -/-! ### Intervals as finsets -/ - -namespace finset -section preorder -variables {α : Type*} [decidable_eq α] [preorder α] [locally_finite_order α] - -/-- The finset of elements `x` such that `a ≤ x ∧ x ≤ b`. Basically `set.Icc a b` as a finset. -/ -def Icc (a b : α) : finset α := -locally_finite_order.finset_Icc a b - -/- `Icc` is treated separately from `Ico`, `Ioc`, `Ioo` because defining it doesn't require -`decidable_rel (≤)`. Useful in simple constructions, like `subtype.locally_finite_order`. -/ -@[simp, norm_cast] lemma coe_Icc (a b : α) : (Icc a b : set α) = set.Icc a b := -locally_finite_order.coe_finset_Icc a b - -@[simp] lemma mem_Icc {a b x : α} : x ∈ Icc a b ↔ a ≤ x ∧ x ≤ b := -by rw [←set.mem_Icc, ←coe_Icc, mem_coe] - -variable [decidable_rel ((≤) : α → α → Prop)] - -/-- The finset of elements `x` such that `a ≤ x ∧ x < b`. Basically `set.Ico a b` as a finset. -/ -def Ico (a b : α) : finset α := -(Icc a b).filter (λ x, ¬b ≤ x) - -/-- The finset of elements `x` such that `a < x ∧ x ≤ b`. Basically `set.Ioc a b` as a finset. -/ -def Ioc (a b : α) : finset α := -(Icc a b).filter (λ x, ¬x ≤ a) - -/-- The finset of elements `x` such that `a < x ∧ x < b`. Basically `set.Ioo a b` as a finset. -/ -def Ioo (a b : α) : finset α := -(Ico a b).filter (λ x, ¬x ≤ a) - -@[simp, norm_cast] lemma coe_Ico (a b : α) : (Ico a b : set α) = set.Ico a b := -by { ext x, rw [mem_coe, Ico, mem_filter, mem_Icc, and_assoc, ←lt_iff_le_not_le, set.mem_Ico] } - -@[simp] lemma mem_Ico {a b x : α} : x ∈ Ico a b ↔ a ≤ x ∧ x < b := -by rw [←set.mem_Ico, ←coe_Ico, mem_coe] - -@[simp, norm_cast] lemma coe_Ioc (a b : α) : (Ioc a b : set α) = set.Ioc a b := -by { ext x, rw [mem_coe, Ioc, mem_filter, mem_Icc, and.right_comm, ←lt_iff_le_not_le, set.mem_Ioc] } - -@[simp] lemma mem_Ioc {a b x : α} : x ∈ Ioc a b ↔ a < x ∧ x ≤ b := -by rw [←set.mem_Ioc, ←coe_Ioc, mem_coe] - -@[simp, norm_cast] lemma coe_Ioo (a b : α) : (Ioo a b : set α) = set.Ioo a b := -by { ext x, rw [mem_coe, Ioo, mem_filter, mem_Ico, and.right_comm, ←lt_iff_le_not_le, set.mem_Ioo] } - -@[simp] lemma mem_Ioo {a b x : α} : x ∈ Ioo a b ↔ a < x ∧ x < b := -by rw [←set.mem_Ioo, ←coe_Ioo, mem_coe] - -theorem Ico_subset_Ico {a₁ b₁ a₂ b₂ : α} (ha : a₂ ≤ a₁) (hb : b₁ ≤ b₂) : - Ico a₁ b₁ ⊆ Ico a₂ b₂ := -begin - rintro x hx, - rw mem_Ico at ⊢ hx, - exact ⟨ha.trans hx.1, hx.2.trans_le hb⟩, -end - -end preorder - -section partial_order -variables {α : Type*} [decidable_eq α] [partial_order α] [decidable_rel ((≤) : α → α → Prop)] - [locally_finite_order α] - -end partial_order - -section order_top -variables {α : Type*} [decidable_eq α] [order_top α] [locally_finite_order α] - -/-- The finset of elements `x` such that `a ≤ x`. Basically `set.Ici a` as a finset. -/ -def Ici (a : α) : finset α := -Icc a ⊤ - -@[simp, norm_cast] lemma coe_Ici (a : α) : (Ici a : set α) = set.Ici a := -by rw [Ici, coe_Icc, set.Icc_top] - -@[simp] lemma mem_Ici {a x : α} : x ∈ Ici a ↔ a ≤ x := -by rw [←set.mem_Ici, ←coe_Ici, mem_coe] - -variable [decidable_rel ((≤) : α → α → Prop)] - -/-- The finset of elements `x` such that `a < x`. Basically `set.Ioi a` as a finset. -/ -def Ioi (a : α) : finset α := -Ioc a ⊤ - -@[simp, norm_cast] lemma coe_Ioi (a : α) : (Ioi a : set α) = set.Ioi a := -by rw [Ioi, coe_Ioc, set.Ioc_top] - -@[simp] lemma mem_Ioi {a x : α} : x ∈ Ioi a ↔ a < x := -by rw [←set.mem_Ioi, ←coe_Ioi, mem_coe] - -end order_top - -section order_bot -variables {α : Type*} [decidable_eq α] [order_bot α] [locally_finite_order α] - -/-- The finset of elements `x` such that `x ≤ b`. Basically `set.Iic b` as a finset. -/ -def Iic (b : α) : finset α := -Icc ⊥ b - -@[simp, norm_cast] lemma coe_Iic (b : α) : (Iic b : set α) = set.Iic b := -by rw [Iic, coe_Icc, set.Icc_bot] - -@[simp] lemma mem_Iic {b x : α} : x ∈ Iic b ↔ x ≤ b := -by rw [←set.mem_Iic, ←coe_Iic, mem_coe] - -variable [decidable_rel ((≤) : α → α → Prop)] - -/-- The finset of elements `x` such that `x < b`. Basically `set.Iio b` as a finset. -/ -def Iio (b : α) : finset α := -Ico ⊥ b - -@[simp, norm_cast] lemma coe_Iio (b : α) : (Iio b : set α) = set.Iio b := -by rw [Iio, coe_Ico, set.Ico_bot] - -@[simp] lemma mem_Iio {b x : α} : x ∈ Iio b ↔ x < b := -by rw [←set.mem_Iio, ←coe_Iio, mem_coe] - -end order_bot +open finset nat section ordered_cancel_add_comm_monoid -variables {α : Type*} [decidable_eq α] [ordered_cancel_add_comm_monoid α] [has_exists_add_of_le α] - [locally_finite_order α] +variables {α : Type*} [ordered_cancel_add_comm_monoid α] [has_exists_add_of_le α] [has_sub α] + [has_ordered_sub α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] lemma image_add_const_Icc (a b c : α) : (Icc a b).image ((+) c) = Icc (a + c) (b + c) := begin @@ -156,8 +36,6 @@ begin exact ⟨a + y, mem_Icc.2 ⟨le_of_add_le_add_right hx.1, le_of_add_le_add_right hx.2⟩, hy⟩, end -variable [decidable_rel ((≤) : α → α → Prop)] - lemma image_add_const_Ico (a b c : α) : (Ico a b).image ((+) c) = Ico (a + c) (b + c) := begin ext x, @@ -391,27 +269,4 @@ begin { simp [range_eq_Ico, Ico.image_const_sub] } end --- TODO We don't yet attempt to reproduce the entire interface for `Ico` for `Ico_ℤ`. - -/-- `Ico_ℤ l u` is the set of integers `l ≤ k < u`. -/ -def Ico_ℤ (l u : ℤ) : finset ℤ := -(finset.range (u - l).to_nat).map - { to_fun := λ n, n + l, - inj' := λ n m h, by simpa using h } - -@[simp] lemma Ico_ℤ.mem {n m l : ℤ} : l ∈ Ico_ℤ n m ↔ n ≤ l ∧ l < m := -begin - dsimp [Ico_ℤ], - simp only [int.lt_to_nat, exists_prop, mem_range, add_comm, function.embedding.coe_fn_mk, - mem_map], - split, - { rintro ⟨a, ⟨h, rfl⟩⟩, - exact ⟨int.le.intro rfl, lt_sub_iff_add_lt'.mp h⟩ }, - { rintro ⟨h₁, h₂⟩, - use (l - n).to_nat, - split; simp [h₁, h₂], } -end - -@[simp] lemma Ico_ℤ.card (l u : ℤ) : (Ico_ℤ l u).card = (u - l).to_nat := by simp [Ico_ℤ] - end finset diff --git a/src/data/fintype/intervals.lean b/src/data/fintype/intervals.lean deleted file mode 100644 index 31bea8ad0ecde..0000000000000 --- a/src/data/fintype/intervals.lean +++ /dev/null @@ -1,76 +0,0 @@ -/- -Copyright (c) 2019 Scott Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison --/ -import data.set.intervals -import data.set.finite -import data.pnat.intervals - -/-! -# fintype instances for intervals - -We provide `fintype` instances for `Ico l u`, for `l u : ℕ`, and for `l u : ℤ`. --/a < b ↔ a + 1 ≤ b - -namespace set - -instance Ico_ℕ_fintype (l u : ℕ) : fintype (Ico l u) := -fintype.of_finset (finset.Ico l u) $ - (λ n, by { simp only [mem_Ico, finset.Ico.mem], }) - -@[simp] lemma Ico_ℕ_card (l u : ℕ) : fintype.card (Ico l u) = u - l := -calc fintype.card (Ico l u) = (finset.Ico l u).card : fintype.card_of_finset _ _ - ... = u - l : finset.Ico.card l u - -instance Ico_pnat_fintype (l u : ℕ+) : fintype (Ico l u) := -fintype.of_finset (pnat.Ico l u) $ - (λ n, by { simp only [mem_Ico, pnat.Ico.mem], }) - -@[simp] lemma Ico_pnat_card (l u : ℕ+) : fintype.card (Ico l u) = u - l := -calc fintype.card (Ico l u) = (pnat.Ico l u).card : fintype.card_of_finset _ _ - ... = u - l : pnat.Ico.card l u - -instance Ico_ℤ_fintype (l u : ℤ) : fintype (Ico l u) := -fintype.of_finset (finset.Ico_ℤ l u) $ - (λ n, by { simp only [mem_Ico, finset.Ico_ℤ.mem], }) - -instance Ioo_ℤ_fintype (l u : ℤ) : fintype (Ioo l u) := -fintype.of_finset (finset.Ico_ℤ (l+1) u) $ λ n, - by simp only [mem_Ioo, finset.Ico_ℤ.mem, int.add_one_le_iff, iff_self, implies_true_iff] - -instance Icc_ℤ_fintype (l u : ℤ) : fintype (Icc l u) := -fintype.of_finset (finset.Ico_ℤ l (u+1)) $ λ n, - by simp only [mem_Icc, finset.Ico_ℤ.mem, int.lt_add_one_iff, iff_self, implies_true_iff] - -instance Ioc_ℤ_fintype (l u : ℤ) : fintype (Ioc l u) := -fintype.of_finset (finset.Ico_ℤ (l+1) (u+1)) $ λ n, - by simp only [mem_Ioc, finset.Ico_ℤ.mem, int.add_one_le_iff, int.lt_add_one_iff, - iff_self, implies_true_iff] - -lemma Ico_ℤ_finite (l u : ℤ) : set.finite (Ico l u) := ⟨set.Ico_ℤ_fintype l u⟩ -lemma Ioo_ℤ_finite (l u : ℤ) : set.finite (Ioo l u) := ⟨set.Ioo_ℤ_fintype l u⟩ -lemma Icc_ℤ_finite (l u : ℤ) : set.finite (Icc l u) := ⟨set.Icc_ℤ_fintype l u⟩ -lemma Ioc_ℤ_finite (l u : ℤ) : set.finite (Ioc l u) := ⟨set.Ioc_ℤ_fintype l u⟩ - -@[simp] lemma Ico_ℤ_card (l u : ℤ) : fintype.card (Ico l u) = (u - l).to_nat := -calc fintype.card (Ico l u) = (finset.Ico_ℤ l u).card : fintype.card_of_finset _ _ - ... = (u - l).to_nat : finset.Ico_ℤ.card l u - -@[simp] lemma Ioo_ℤ_card (l u : ℤ) : fintype.card (Ioo l u) = (u - l - 1).to_nat := -calc fintype.card (Ioo l u) = (finset.Ico_ℤ (l+1) u).card : fintype.card_of_finset _ _ - ... = (u - (l+1)).to_nat : finset.Ico_ℤ.card (l+1) u - ... = (u - l - 1).to_nat : congr_arg _ $ sub_add_eq_sub_sub u l 1 - -@[simp] lemma Icc_ℤ_card (l u : ℤ) : fintype.card (Icc l u) = (u + 1 - l).to_nat := -calc fintype.card (Icc l u) = (finset.Ico_ℤ l (u+1)).card : fintype.card_of_finset _ _ - ... = (u + 1 - l).to_nat : finset.Ico_ℤ.card l (u+1) - -@[simp] lemma Ioc_ℤ_card (l u : ℤ) : fintype.card (Ioc l u) = (u - l).to_nat := -calc fintype.card (Ioc l u) = (finset.Ico_ℤ (l+1) (u+1)).card : fintype.card_of_finset _ _ - ... = ((u+1) - (l+1)).to_nat : finset.Ico_ℤ.card (l+1) (u+1) - ... = (u - l).to_nat : congr_arg _ $ add_sub_add_right_eq_sub u l 1 - --- TODO other useful instances: fin n, zmod? - -end set diff --git a/src/data/int/intervals.lean b/src/data/int/intervals.lean new file mode 100644 index 0000000000000..7c4dcb6f54ad8 --- /dev/null +++ b/src/data/int/intervals.lean @@ -0,0 +1,79 @@ +/- +Copyright (c) 2021 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import data.int.basic +import order.locally_finite + +/-! +# Intervals of integers + +This file proves that `ℤ` is a `locally_finite_order` and calculates the cardinality of its +intervals as finsets and fintypes. +-/ + +open finset int + +instance : locally_finite_order ℤ := +{ finset_Icc := λ a b, (Iio (b + 1 - a).to_nat).map ⟨λ n, n + a, λ _, by simp only + [imp_self, forall_const, add_left_inj, int.coe_nat_inj']⟩, + finset_mem_Icc := λ a b x, begin + rw [mem_map, set.mem_Icc, function.embedding.coe_fn_mk], + split, + { rintro ⟨n, hn, hx⟩, + rw [←hx, le_add_iff_nonneg_left], + rw [mem_Iio_iff, int.lt_to_nat, lt_sub_iff_add_lt, int.lt_add_one_iff] at hn, + exact ⟨int.coe_nat_nonneg _, hn⟩ }, + rintro h, + refine ⟨(x - a).to_nat, _, by rw [int.to_nat_sub_of_le h.1, int.sub_add_cancel]⟩, + rw mem_Iio_iff, + have hb := int.lt_add_one_of_le h.2, + exact (int.to_nat_lt_to_nat (sub_pos.2 (h.1.trans_lt hb))).2 (sub_lt_sub_right hb _), + end } + +namespace int +variables (a b : ℤ) + +/-- `int_Ico a b` is the set of integers `b ≤ k < a`. -/ +def int_Ico (a b : ℤ) : finset ℤ := +(finset.range (b - a).to_nat).map + { to_fun := λ n, n + b, + inj' := λ n m h, by simpa using h } + +@[simp] lemma int_Ico.mem {n m l : ℤ} : l ∈ int_Ico n m ↔ n ≤ l ∧ l < m := +begin + dsimp [int_Ico], + simp only [int.lt_to_nat, exists_prop, mem_range, add_comm, function.embedding.coe_fn_mk, + mem_map], + split, + { rintro ⟨a, ⟨h, rfl⟩⟩, + exact ⟨int.le.intro rfl, lt_sub_iff_add_lt'.mp h⟩ }, + { rintro ⟨h₁, h₂⟩, + use (l - n).to_nat, + split; simp [h₁, h₂], } +end + +@[simp] lemma int_Ico.card (a b : ℤ) : (int_Ico a b).card = (b - a).to_nat := by simp [int_Ico] + +@[simp] lemma card_finset_Icc : (Icc a b).card = (b + 1 - a).to_nat := sorry + +@[simp] lemma card_finset_Ico : (Ico a b).card = (b - a).to_nat := sorry + +@[simp] lemma card_finset_Ioc : (Ioc a b).card = (b - a).to_nat := sorry + +@[simp] lemma card_finset_Ioo : (Ioo a b).card = (b - a - 1).to_nat := sorry + +@[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = (b + 1 - a).to_nat := +by rw [←card_finset_Icc, fintype.card_of_finset] + +@[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = (b - a).to_nat := +by rw [←card_finset_Ico, fintype.card_of_finset] + +@[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = (b - a).to_nat := +by rw [←card_finset_Ioc, fintype.card_of_finset] + +@[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = (b - a - 1).to_nat := +by rw [←card_finset_Ioo, fintype.card_of_finset] + +end int diff --git a/src/data/nat/intervals.lean b/src/data/nat/intervals.lean new file mode 100644 index 0000000000000..f6907f819de87 --- /dev/null +++ b/src/data/nat/intervals.lean @@ -0,0 +1,51 @@ +/- +Copyright (c) 2021 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import data.nat.basic +import order.locally_finite + +/-! +# Intervals of naturals + +This file proves that `ℕ` is a `locally_finite_order` and calculates the cardinality of its +intervals as finsets and fintypes. +-/ + +open finset nat + +instance : locally_finite_order ℕ := +{ finset_Icc := λ a b, (list.range' a (b + 1 - a)).to_finset, + finset_mem_Icc := λ a b x, begin + rw [list.mem_to_finset, list.mem_range'], + cases le_or_lt a b, + { rw [nat.add_sub_cancel' (nat.lt_succ_of_le h).le, nat.lt_succ_iff] }, + { rw [nat.sub_eq_zero_iff_le.2 h, add_zero], + exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)) } +end } + +namespace nat +variables (a b : ℕ) + +@[simp] lemma card_finset_Icc : (Icc a b).card = b + 1 - a := sorry + +@[simp] lemma card_finset_Ico : (Ico a b).card = b - a := sorry + +@[simp] lemma card_finset_Ioc : (Ioc a b).card = b - a := sorry + +@[simp] lemma card_finset_Ioo : (Ioo a b).card = b - a - 1 := sorry + +@[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := +by rw [←card_finset_Icc, fintype.card_of_finset] + +@[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = b - a := +by rw [←card_finset_Ico, fintype.card_of_finset] + +@[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = b - a := +by rw [←card_finset_Ioc, fintype.card_of_finset] + +@[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := +by rw [←card_finset_Ioo, fintype.card_of_finset] + +end nat diff --git a/src/data/pnat/intervals.lean b/src/data/pnat/intervals.lean index ee8b1488dbf75..6f6398b4f9c62 100644 --- a/src/data/pnat/intervals.lean +++ b/src/data/pnat/intervals.lean @@ -1,23 +1,61 @@ /- Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Scott Morrison, Yaël Dillies -/ import data.pnat.basic -import data.finset.intervals +import order.locally_finite + +/-! +# Intervals of positive naturals + +This file proves that `ℕ+` is a `locally_finite_order` and calculates the cardinality of its +intervals as finsets and fintypes. +-/ + +open finset pnat + +instance : locally_finite_order ℕ+ := +{ finset_Icc := λ a b, (list.range' a (b + 1 - a)).to_finset, + finset_mem_Icc := λ a b x, begin + rw [list.mem_to_finset, list.mem_range'], + cases le_or_lt a b, + { rw [nat.add_sub_cancel' (nat.lt_succ_of_le h).le, nat.lt_succ_iff] }, + { rw [nat.sub_eq_zero_iff_le.2 h, add_zero], + exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)) } +end } namespace pnat +variables (a b : ℕ+) -/-- `Ico l u` is the set of positive natural numbers `l ≤ k < u`. -/ -def Ico (l u : ℕ+) : finset ℕ+ := -(finset.Ico l u).attach.map - { to_fun := λ n, ⟨(n : ℕ), lt_of_lt_of_le l.2 (finset.Ico.mem.1 n.2).1⟩, +/-- `Ico a b` is the set of positive natural numbers `b ≤ k < a`. -/ +def pnat_Ico (a b : ℕ+) : finset ℕ+ := +(finset.Ico a b).attach.map + { to_fun := λ n, ⟨(n : ℕ), lt_of_lt_of_le a.2 (finset.Ico.mem.1 n.2).1⟩, -- why can't we do this directly? inj' := λ n m h, subtype.eq (by { replace h := congr_arg subtype.val h, exact h }) } -@[simp] lemma Ico.mem : ∀ {n m l : ℕ+}, l ∈ Ico n m ↔ n ≤ l ∧ l < m := -by { rintro ⟨n, hn⟩ ⟨m, hm⟩ ⟨l, hl⟩, simp [pnat.Ico] } +@[simp] lemma pnat_Ico.mem : ∀ {n m l : ℕ+}, l ∈ pnat_Ico n m ↔ n ≤ l ∧ l < m := +by { rintro ⟨n, hn⟩ ⟨m, hm⟩ ⟨l, hl⟩, simp [pnat_Ico] } + +@[simp] lemma card_finset_Icc : (Icc a b).card = b + 1 - a := sorry + +@[simp] lemma card_finset_Ico : (Ico a b).card = b - a := sorry + +@[simp] lemma card_finset_Ioc : (Ioc a b).card = b - a := sorry + +@[simp] lemma card_finset_Ioo : (Ioo a b).card = b - a - 1 := sorry + +@[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := +by rw [←card_finset_Icc, fintype.card_of_finset] + +@[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = b - a := +by rw [←card_finset_Ico, fintype.card_of_finset] + +@[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = b - a := +by rw [←card_finset_Ioc, fintype.card_of_finset] -@[simp] lemma Ico.card (l u : ℕ+) : (Ico l u).card = u - l := by simp [pnat.Ico] +@[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := +by rw [←card_finset_Ioo, fintype.card_of_finset] end pnat diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 6fdad6fa19df6..c62ed005b00df 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -8,7 +8,6 @@ import data.finset.preimage import data.finset.sort import data.set.finite import data.set.intervals.basic -import data.int.basic /-! # Locally finite orders @@ -23,11 +22,26 @@ Further, if the order is bounded above (resp. below), then we can also make sens ## Main declarations In a `locally_finite_order`, -* `finset.` `Icc`/`Ico`/`Ioc`/`Ioo`: Closed/open-closed/open interval as a finset. -* `finset.` `Ici`/`Ioi`: Closed/open-infinite interval as a finset if the order is also an - `order_top`. -* `finset.` `Iic`/`Iio`: Infinite-closed/open interval as a finset if the order is also an - `order_bot`. +* `finset.Icc`: Closed-closed interval as a finset. +* `finset.Ico`: Closed-open interval as a finset. +* `finset.Ioc`: Open-closed interval as a finset. +* `finset.Ioo`: Open-open interval as a finset. +* `multiset.Icc`: Closed-closed interval as a multiset. +* `multiset.Ico`: Closed-open interval as a multiset. +* `multiset.Ioc`: Open-closed interval as a multiset. +* `multiset.Ioo`: Open-open interval as a finset. + +When it's also an `order_top`, +* `finset.Ici`: Closed-infinite interval as a finset. +* `finset.Ioi`: Open-infinite interval as a finset. +* `multiset.Ici`: Closed-infinite interval as a multiset. +* `multiset.Ioi`: Open-infinite interval as a multiset. + +When it's also an `order_bot`, +* `finset.Iic`: Infinite-open interval as a finset. +* `finset.Iio`: Infinite-closed interval as a finset. +* `multiset.Iic`: Infinite-open interval as a multiset. +* `multiset.Iio`: Infinite-closed interval as a multiset. ## Instances @@ -42,8 +56,8 @@ We provide a `locally_finite_order` instance for TODO: All the API in `set.intervals.basic` could be copied here. But this is unfortunately labor- and maintenance-intensive. -TODO:nce we have the `has_enum` typeclass (any non-top element has a least greater element or any -non-bot element has a greatest lesser element), we can provide an `has_enum` typeclass using +TODO: Once we have the `succ_order` (any non-top element has a least greater element) typeclass, we +can provide `succ_order α` from `linear_order α` and `locally_finite_order α` using ```lean lemma exists_min_greater [linear_order α] [locally_finite_order α] {x ub : α} (hx : x < ub) : @@ -63,79 +77,345 @@ end ``` Note that the converse is not true. Consider `{-2^z | z : ℤ} ∪ {2^z | z : ℤ}`. Any element has a -least greater element and a greatest lesser element, so it `has_enum` (both ways!), but it's not -locally finite as `Icc (-1) 1` is infinite. +successor, so it is a `succ_order`, but it's not locally finite as `Icc (-1) 1` is infinite. TODO: From `linear_order α`, `no_top_order α`, `locally_finite_order α`, we can also define an order isomorphism `α ≃ ℕ` or `α ≃ ℤ`, depending on whether we have `order_bot α` or `no_bot_order α` and `nonempty α`. When `order_bot α`, we can match `a : α` to `(Iio a).card`. -/ +lemma and_and_and_comm (a b c d : Prop) : (a ∧ b) ∧ c ∧ d ↔ (a ∧ c) ∧ b ∧ d := +by rw [←and_assoc, @and.right_comm a, and_assoc] + +section tactic +open tactic + +meta def finset_Ico_laws_tac : tactic unit := +whnf_target >> intros >> to_expr``(begin + unfold finset_Ico, + rw [finset.mem_filter, finset_mem_Icc, and_assoc, lt_iff_le_not_le] +end) >>= exact + +meta def finset_Ioc_laws_tac : tactic unit := +whnf_target >> intros >> to_expr``(by { unfold locally_finite_order.finset_Ioc, + rw [finset.mem_filter, finset_mem_Icc, and.right_comm, lt_iff_le_not_le] }) >>= exact + +meta def finset_Ioo_laws_tac : tactic unit := +whnf_target >> intros >> to_expr``(by { unfold locally_finite_order.finset_Ioo, + rw [finset.mem_filter, finset_mem_Icc, and_and_and_comm, lt_iff_le_not_le, lt_iff_le_not_le] }) + >>= exact + +end tactic + open finset -class locally_finite_order (α : Type*) [preorder α] := +class locally_finite_order (α : Type*) [preorder α] [decidable_rel ((≤) : α → α → Prop)] := (finset_Icc : α → α → finset α) -(coe_finset_Icc : ∀ a b : α, (finset_Icc a b : set α) = set.Icc a b) +(finset_Ico : α → α → finset α := λ a b, (finset_Icc a b).filter (λ x, ¬b ≤ x)) +(finset_Ioc : α → α → finset α := λ a b, (finset_Icc a b).filter (λ x, ¬x ≤ a)) +(finset_Ioo : α → α → finset α := λ a b, (finset_Icc a b).filter (λ x, ¬x ≤ a ∧ ¬b ≤ x)) +(finset_mem_Icc : ∀ a b x : α, x ∈ finset_Icc a b ↔ a ≤ x ∧ x ≤ b) +(finset_mem_Ico : ∀ a b x : α, x ∈ finset_Ico a b ↔ a ≤ x ∧ x < b . finset_Ico_laws_tac) +(finset_mem_Ioc : ∀ a b x : α, x ∈ finset_Ioc a b ↔ a < x ∧ x ≤ b . finset_Ioc_laws_tac) +(finset_mem_Ioo : ∀ a b x : α, x ∈ finset_Ioo a b ↔ a < x ∧ x < b . finset_Ioo_laws_tac) -namespace set -variables {α : Type*} [preorder α] [locally_finite_order α] +variables {α β : Type*} + +/-! ### Intervals as finsets -/ + +namespace finset +section preorder +variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] + +/-- The finset of elements `x` such that `a ≤ x` and `x ≤ b`. Basically `set.Icc a b` as a finset. +-/ +def Icc (a b : α) : finset α := +locally_finite_order.finset_Icc a b + +/-- The finset of elements `x` such that `a ≤ x` and `x < b`. Basically `set.Ico a b` as a finset. +-/ +def Ico (a b : α) : finset α := +locally_finite_order.finset_Ico a b + +/-- The finset of elements `x` such that `a < x` and `x ≤ b`. Basically `set.Ioc a b` as a finset. +-/ +def Ioc (a b : α) : finset α := +locally_finite_order.finset_Ioc a b + +/-- The finset of elements `x` such that `a < x` and `x < b`. Basically `set.Ioo a b` as a finset. +-/ +def Ioo (a b : α) : finset α := +locally_finite_order.finset_Ioo a b + +@[simp] lemma mem_Icc {a b x : α} : x ∈ Icc a b ↔ a ≤ x ∧ x ≤ b := +locally_finite_order.finset_mem_Icc a b x + +@[simp] lemma mem_Ico {a b x : α} : x ∈ Ico a b ↔ a ≤ x ∧ x < b := +locally_finite_order.finset_mem_Ico a b x + +@[simp] lemma mem_Ioc {a b x : α} : x ∈ Ioc a b ↔ a < x ∧ x ≤ b := +locally_finite_order.finset_mem_Ioc a b x + +@[simp] lemma mem_Ioo {a b x : α} : x ∈ Ioo a b ↔ a < x ∧ x < b := +locally_finite_order.finset_mem_Ioo a b x + +@[simp, norm_cast] lemma coe_Icc (a b : α) : (Icc a b : set α) = set.Icc a b := +by { ext, rw [mem_coe, mem_Icc, set.mem_Icc] } -lemma Icc_finite (a b : α) : (Icc a b).finite := +@[simp, norm_cast] lemma coe_Ico (a b : α) : (Ico a b : set α) = set.Ico a b := +by { ext, rw [mem_coe, mem_Ico, set.mem_Ico] } + +@[simp, norm_cast] lemma coe_Ioc (a b : α) : (Ioc a b : set α) = set.Ioc a b := +by { ext, rw [mem_coe, mem_Ioc, set.mem_Ioc] } + +@[simp, norm_cast] lemma coe_Ioo (a b : α) : (Ioo a b : set α) = set.Ioo a b := +by { ext, rw [mem_coe, mem_Ioo, set.mem_Ioo] } + +theorem Ico_subset_Ico {a₁ b₁ a₂ b₂ : α} (ha : a₂ ≤ a₁) (hb : b₁ ≤ b₂) : + Ico a₁ b₁ ⊆ Ico a₂ b₂ := begin - rw ←locally_finite_order.coe_finset_Icc, - exact (locally_finite_order.finset_Icc a b).finite_to_set, + rintro x hx, + rw mem_Ico at ⊢ hx, + exact ⟨ha.trans hx.1, hx.2.trans_le hb⟩, end -lemma Ico_finite (a b : α) : (Ico a b).finite := -(Icc_finite a b).subset Ico_subset_Icc_self +end preorder -lemma Ioc_finite (a b : α) : (Ioc a b).finite := -(Icc_finite a b).subset Ioc_subset_Icc_self +section partial_order +variables [partial_order α] [decidable_rel ((≤) : α → α → Prop)] + [locally_finite_order α] -lemma Ioo_finite (a b : α) : (Ioo a b).finite := -(Icc_finite a b).subset Ioo_subset_Icc_self +end partial_order -end set +section order_top +variables [order_top α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] -open finset +/-- The finset of elements `x` such that `a ≤ x`. Basically `set.Ici a` as a finset. -/ +def Ici (a : α) : finset α := +Icc a ⊤ + +/-- The finset of elements `x` such that `a < x`. Basically `set.Ioi a` as a finset. -/ +def Ioi (a : α) : finset α := +Ioc a ⊤ + +@[simp, norm_cast] lemma coe_Ici (a : α) : (Ici a : set α) = set.Ici a := +by rw [Ici, coe_Icc, set.Icc_top] + +@[simp, norm_cast] lemma coe_Ioi (a : α) : (Ioi a : set α) = set.Ioi a := +by rw [Ioi, coe_Ioc, set.Ioc_top] + +@[simp] lemma mem_Ici {a x : α} : x ∈ Ici a ↔ a ≤ x := +by rw [←set.mem_Ici, ←coe_Ici, mem_coe] + +@[simp] lemma mem_Ioi {a x : α} : x ∈ Ioi a ↔ a < x := +by rw [←set.mem_Ioi, ←coe_Ioi, mem_coe] + +end order_top + +section order_bot +variables [order_bot α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] + +/-- The finset of elements `x` such that `x ≤ b`. Basically `set.Iic b` as a finset. -/ +def Iic (b : α) : finset α := +Icc ⊥ b + +/-- The finset of elements `x` such that `x < b`. Basically `set.Iio b` as a finset. -/ +def Iio (b : α) : finset α := +Ico ⊥ b + +@[simp, norm_cast] lemma coe_Iic (b : α) : (Iic b : set α) = set.Iic b := +by rw [Iic, coe_Icc, set.Icc_bot] + +@[simp, norm_cast] lemma coe_Iio (b : α) : (Iio b : set α) = set.Iio b := +by rw [Iio, coe_Ico, set.Ico_bot] + +@[simp] lemma mem_Iic {b x : α} : x ∈ Iic b ↔ x ≤ b := +by rw [←set.mem_Iic, ←coe_Iic, mem_coe] + +@[simp] lemma mem_Iio {b x : α} : x ∈ Iio b ↔ x < b := +by rw [←set.mem_Iio, ←coe_Iio, mem_coe] + +end order_bot +end finset + +/-! ### Intervals as multisets -/ + +namespace multiset +section preorder +variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] + +/-- The multiset of elements `x` such that `a ≤ x` and `x ≤ b`. Basically `set.Icc a b` as a +multiset. -/ +def Icc (a b : α) : multiset α := +(finset.Icc a b).val + +-- TODO@Yaël: Nuke `data.multiset.intervals` and redefine `multiset.Ico` here + +/-- The multiset of elements `x` such that `a < x` and `x ≤ b`. Basically `set.Ioc a b` as a +multiset. -/ +def Ioc (a b : α) : multiset α := +(finset.Ioc a b).val + +/-- The multiset of elements `x` such that `a < x` and `x < b`. Basically `set.Ioo a b` as a +multiset. -/ +def Ioo (a b : α) : multiset α := +(finset.Ioo a b).val + +@[simp] lemma mem_Icc {a b x : α} : x ∈ Icc a b ↔ a ≤ x ∧ x ≤ b := +by rw [Icc, ←finset.mem_def, finset.mem_Icc] + +@[simp] lemma mem_Ioc {a b x : α} : x ∈ Ioc a b ↔ a < x ∧ x ≤ b := +by rw [Ioc, ←finset.mem_def, finset.mem_Ioc] + +@[simp] lemma mem_Ioo {a b x : α} : x ∈ Ioo a b ↔ a < x ∧ x < b := +by rw [Ioo, ←finset.mem_def, finset.mem_Ioo] + +end preorder + +section order_top +variables [order_top α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] + +/-- The multiset of elements `x` such that `a ≤ x`. Basically `set.Ici a` as a multiset. -/ +def Ici (a : α) : multiset α := +(finset.Ici a).val + +/-- The multiset of elements `x` such that `a < x`. Basically `set.Ioi a` as a multiset. -/ +def Ioi (a : α) : multiset α := +(finset.Ioi a).val + +@[simp] lemma mem_Ici {a x : α} : x ∈ Ici a ↔ a ≤ x := +by rw [Ici, ←finset.mem_def, finset.mem_Ici] + +@[simp] lemma mem_Ioi {a x : α} : x ∈ Ioi a ↔ a < x := +by rw [Ioi, ←finset.mem_def, finset.mem_Ioi] + +end order_top + +section order_bot +variables [order_bot α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] + +/-- The multiset of elements `x` such that `x ≤ b`. Basically `set.Iic b` as a multiset. -/ +def Iic (b : α) : multiset α := +(finset.Iic b).val + +/-- The multiset of elements `x` such that `x < b`. Basically `set.Iio b` as a multiset. -/ +def Iio (b : α) : multiset α := +(finset.Iio b).val + +@[simp] lemma mem_Iic {b x : α} : x ∈ Iic b ↔ x ≤ b := +by rw [Iic, ←finset.mem_def, finset.mem_Iic] + +@[simp] lemma mem_Iio {b x : α} : x ∈ Iio b ↔ x < b := +by rw [Iio, ←finset.mem_def, finset.mem_Iio] + +end order_bot +end multiset + +/-! ### Finiteness of `set` intervals -/ + +namespace set +section preorder +variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] (a b : α) + +instance fintype_Icc : fintype (Icc a b) := +fintype.of_finset (finset.Icc a b) (λ x, by rw [finset.mem_Icc, mem_Icc]) + +instance fintype_Ico : fintype (Ico a b) := +fintype.of_finset (finset.Ico a b) (λ x, by rw [finset.mem_Ico, mem_Ico]) + +instance fintype_Ioc : fintype (Ioc a b) := +fintype.of_finset (finset.Ioc a b) (λ x, by rw [finset.mem_Ioc, mem_Ioc]) + +instance fintype_Ioo : fintype (Ioo a b) := +fintype.of_finset (finset.Ioo a b) (λ x, by rw [finset.mem_Ioo, mem_Ioo]) + +lemma finite_Icc : (Icc a b).finite := +⟨set.fintype_Icc a b⟩ + +lemma finite_Ico : (Ico a b).finite := +⟨set.fintype_Ico a b⟩ + +lemma finite_Ioc : (Ioc a b).finite := +⟨set.fintype_Ioc a b⟩ + +lemma finite_Ioo : (Ioo a b).finite := +⟨set.fintype_Ioo a b⟩ + +end preorder + +section order_top +variables [order_top α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] (a : α) + +instance fintype_Ici : fintype (Ici a) := +fintype.of_finset (finset.Ici a) (λ x, by rw [finset.mem_Ici, mem_Ici]) + +instance fintype_Ioi : fintype (Ioi a) := +fintype.of_finset (finset.Ioi a) (λ x, by rw [finset.mem_Ioi, mem_Ioi]) + +lemma finite_Ici : (Ici a).finite := +⟨set.fintype_Ici a⟩ + +lemma finite_Ioi : (Ioi a).finite := +⟨set.fintype_Ioi a⟩ + +end order_top + +section order_bot +variables [order_bot α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] (b : α) + +instance fintype_Iic : fintype (Iic b) := +fintype.of_finset (finset.Iic b) (λ x, by rw [finset.mem_Iic, mem_Iic]) + +instance fintype_Iio : fintype (Iio b) := +fintype.of_finset (finset.Iio b) (λ x, by rw [finset.mem_Iio, mem_Iio]) + +lemma finite_Iic : (Iic b).finite := +⟨set.fintype_Iic b⟩ + +lemma finite_Iio : (Iio b).finite := +⟨set.fintype_Iio b⟩ + +end order_bot + +end set /-! ### Instances -/ -noncomputable def locally_finite_order_of_finite_Icc {α : Type*} [preorder α] - (h : ∀ x y : α, (set.Icc x y).finite) : +section +variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] + +noncomputable def locally_finite_order_of_finite_Icc + (h : ∀ a b : α, (set.Icc a b).finite) : locally_finite_order α := -{ finset_Icc := λ x y, (h x y).to_finset, - coe_finset_Icc := λ x y, (h x y).coe_to_finset } +{ finset_Icc := λ a b, (h a b).to_finset, + finset_mem_Icc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Icc], + finset_mem_Ico := λ a b x, begin + rw [finset.mem_filter, finset_mem_Icc, and_assoc, lt_iff_le_not_le], + end } @[priority 900] -noncomputable instance fintype.locally_finite_order {α : Type*} [preorder α] [fintype α] : +noncomputable instance fintype.locally_finite_order [fintype α] : locally_finite_order α := { finset_Icc := λ a b, (set.finite.of_fintype (set.Icc a b)).to_finset, - coe_finset_Icc := λ a b, set.finite.coe_to_finset _ } + finset_mem_Icc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Icc] } -- Should this be called `order_embedding.locally_finite_order`? -/-- Given an order embedding `α ↪o β`, pulls back to `α` the `locally_finite_order` on `β`. -/ -noncomputable def locally_finite_order.lift {α β : Type*} [partial_order β] [locally_finite_order β] +/-- Given an order embedding `α ↪o β`, pulls back the `locally_finite_order` on `β` to `α`. -/ +noncomputable def locally_finite_order.lift {β : Type*} [partial_order β] [locally_finite_order β] [decidable_eq β] [partial_order α] [decidable_eq α] (f : α ↪o β) : locally_finite_order α := { finset_Icc := λ a b, (Icc (f a) (f b)).preimage f (f.to_embedding.injective.inj_on _), - coe_finset_Icc := λ a b, begin - ext, - simp only [coe_Icc, coe_preimage, iff_self, set.mem_preimage, set.mem_Icc, - order_embedding.le_iff_le], - end } + finset_mem_Icc := λ a b x, by rw [mem_preimage, mem_Icc, f.le_iff_le, f.le_iff_le] } -instance subtype.locally_finite_order {α : Type*} [preorder α] [locally_finite_order α] - [decidable_eq α] {p : α → Prop} [decidable_pred p] : +instance subtype.locally_finite_order [locally_finite_order α] {p : α → Prop} [decidable_pred p] : locally_finite_order (subtype p) := begin haveI : preorder (subtype p) := by apply_instance, - haveI : decidable_eq (subtype p) := by apply_instance, + haveI : decidable_rel ((≤) : subtype p → subtype p → Prop) := by apply_instance, exact { finset_Icc := λ a b, finset.subtype p (Icc a.val b.val), - coe_finset_Icc := λ a b, begin - ext, - rw [finset.subtype, coe_map, set.mem_image, set.mem_Icc], + finset_mem_Icc := λ a b x, begin + rw [finset.subtype, mem_map, set.mem_image, set.mem_Icc], dsimp, split, { rintro ⟨⟨y, hy⟩, -, h⟩, @@ -150,31 +430,4 @@ begin end }, end -instance nat.locally_finite_order : locally_finite_order ℕ := -{ finset_Icc := λ a b, (list.range' a (b + 1 - a)).to_finset, - coe_finset_Icc := λ a b, begin - ext, - rw [mem_coe, list.mem_to_finset, list.mem_range', set.mem_Icc], - cases le_or_lt a b, - { rw [nat.add_sub_cancel' (nat.lt_succ_of_le h).le, nat.lt_succ_iff] }, - rw [nat.sub_eq_zero_iff_le.2 h, add_zero], - exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)), -end } - -instance int.locally_finite_order : locally_finite_order ℤ := -{ finset_Icc := λ a b, (Iio (b + 1 - a).to_nat).map ⟨λ n, n + a, λ _, by simp only - [imp_self, forall_const, add_left_inj, int.coe_nat_inj']⟩, - coe_finset_Icc := λ a b, begin - ext, - rw [mem_coe, mem_map, set.mem_Icc, function.embedding.coe_fn_mk], - split, - { rintro ⟨n, hn, hx⟩, - rw [←hx, le_add_iff_nonneg_left], - rw [mem_Iio_iff, int.lt_to_nat, lt_sub_iff_add_lt, int.lt_add_one_iff] at hn, - exact ⟨int.coe_nat_nonneg _, hn⟩ }, - rintro h, - refine ⟨(x - a).to_nat, _, by rw [int.to_nat_sub_of_le h.1, int.sub_add_cancel]⟩, - rw mem_Iio_iff, - have hb := int.lt_add_one_of_le h.2, - exact (int.to_nat_lt_to_nat (sub_pos.2 (h.1.trans_lt hb))).2 (sub_lt_sub_right hb _), - end } +end From 86171c58b1c3bfcfb24eed7712d1b56c014b2267 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Tue, 14 Sep 2021 19:35:05 +0200 Subject: [PATCH 07/41] finish API --- src/algebra/big_operators/intervals.lean | 4 +- src/algebra/geom_sum.lean | 2 +- src/analysis/analytic/composition.lean | 2 +- src/analysis/analytic/inverse.lean | 4 +- src/data/int/intervals.lean | 104 ++++++++++------- src/data/nat/intervals.lean | 50 ++++++++- src/data/nat/multiplicity.lean | 2 +- src/data/nat/totient.lean | 4 +- src/data/pnat/intervals.lean | 57 +++++----- src/data/polynomial/inductions.lean | 2 +- src/order/locally_finite.lean | 135 +++++++++++------------ src/tactic/interval_cases.lean | 2 +- src/topology/metric_space/basic.lean | 2 +- 13 files changed, 216 insertions(+), 154 deletions(-) diff --git a/src/algebra/big_operators/intervals.lean b/src/algebra/big_operators/intervals.lean index a940880bf2397..79b809b4f466f 100644 --- a/src/algebra/big_operators/intervals.lean +++ b/src/algebra/big_operators/intervals.lean @@ -81,7 +81,7 @@ begin (λ (x : Σ (i : ℕ), ℕ) _, (⟨x.2, x.1⟩ : Σ (i : ℕ), ℕ)) _ (λ _ _, rfl) (λ (x : Σ (i : ℕ), ℕ) _, (⟨x.2, x.1⟩ : Σ (i : ℕ), ℕ)) _ (by rintro ⟨⟩ _; refl) (by rintro ⟨⟩ _; refl); - simp only [finset.Ico.mem, sigma.forall, finset.mem_sigma]; + simp only [finset.mem_Ico, sigma.forall, finset.mem_sigma]; rintros a b ⟨⟨h₁,h₂⟩, ⟨h₃, h₄⟩⟩; refine ⟨⟨_, _⟩, ⟨_, _⟩⟩; linarith end @@ -104,7 +104,7 @@ begin cases lt_or_le k m with hkm hkm, { rw [← finset.Ico.image_const_sub (this _ hkm)], refine (prod_image _).symm, - simp only [Ico.mem], + simp only [mem_Ico], rintros i ⟨ki, im⟩ j ⟨kj, jm⟩ Hij, rw [← nat.sub_sub_self (this _ im), Hij, nat.sub_sub_self (this _ jm)] }, { simp [Ico.eq_empty_of_le, nat.sub_le_sub_left, hkm] } diff --git a/src/algebra/geom_sum.lean b/src/algebra/geom_sum.lean index dd792ac2d41c2..8f7f93d90cac9 100644 --- a/src/algebra/geom_sum.lean +++ b/src/algebra/geom_sum.lean @@ -361,7 +361,7 @@ begin = a/b^0 + ∑ (i : ℕ) in Ico 1 n.succ, a/b^i : by rw [pow_zero, nat.div_one] ... = ∑ i in range n.succ, a/b^i : begin rw [range_eq_Ico, ←finset.Ico.insert_succ_bot (nat.succ_pos _), sum_insert], - exact λ h, zero_lt_one.not_le (Ico.mem.1 h).1, + exact λ h, zero_lt_one.not_le (mem_Ico.1 h).1, end ... ≤ a * b/(b - 1) : nat.geom_sum_le hb a _ ... = (a * 1 + a * (b - 1))/(b - 1) diff --git a/src/analysis/analytic/composition.lean b/src/analysis/analytic/composition.lean index a45d17a35e24a..0f0b05283927f 100644 --- a/src/analysis/analytic/composition.lean +++ b/src/analysis/analytic/composition.lean @@ -537,7 +537,7 @@ finset.sigma (finset.Ico m M) (λ (n : ℕ), fintype.pi_finset (λ (i : fin n), @[simp] lemma mem_comp_partial_sum_source_iff (m M N : ℕ) (i : Σ n, (fin n) → ℕ) : i ∈ comp_partial_sum_source m M N ↔ (m ≤ i.1 ∧ i.1 < M) ∧ ∀ (a : fin i.1), 1 ≤ i.2 a ∧ i.2 a < N := -by simp only [comp_partial_sum_source, finset.Ico.mem, fintype.mem_pi_finset, finset.mem_sigma, +by simp only [comp_partial_sum_source, finset.mem_Ico, fintype.mem_pi_finset, finset.mem_sigma, iff_self] /-- Change of variables appearing to compute the composition of partial sums of formal diff --git a/src/analysis/analytic/inverse.lean b/src/analysis/analytic/inverse.lean index 58786d8a47fe9..e696c34b5c6ca 100644 --- a/src/analysis/analytic/inverse.lean +++ b/src/analysis/analytic/inverse.lean @@ -370,7 +370,7 @@ begin refine sum_le_sum_of_subset_of_nonneg _ (λ x hx1 hx2, prod_nonneg (λ j hj, mul_nonneg hr (mul_nonneg (pow_nonneg ha _) (hp _)))), rintros ⟨k, c⟩ hd, - simp only [set.mem_to_finset, Ico.mem, mem_sigma, set.mem_set_of_eq] at hd, + simp only [set.mem_to_finset, mem_Ico, mem_sigma, set.mem_set_of_eq] at hd, simp only [mem_comp_partial_sum_target_iff], refine ⟨hd.2, c.length_le.trans_lt hd.1.2, λ j, _⟩, have : c ≠ composition.single k (zero_lt_two.trans_le hd.1.1), @@ -425,7 +425,7 @@ by simp only [linear_isometry_equiv.norm_map, pow_one, right_inv_coeff_one, begin congr' 1, apply sum_congr rfl (λ j hj, _), - rw [right_inv_coeff _ _ _ (Ico.mem.1 hj).1, norm_neg], + rw [right_inv_coeff _ _ _ (mem_Ico.1 hj).1, norm_neg], end ... ≤ a * ∥(i.symm : F →L[𝕜] E)∥ + ∑ k in Ico 2 (n + 1), a ^ k * (I * (∑ c in ({c | 1 < composition.length c}.to_finset : finset (composition k)), diff --git a/src/data/int/intervals.lean b/src/data/int/intervals.lean index 7c4dcb6f54ad8..a840933c122a7 100644 --- a/src/data/int/intervals.lean +++ b/src/data/int/intervals.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import data.int.basic -import order.locally_finite +import data.nat.intervals /-! # Intervals of integers @@ -16,53 +16,79 @@ intervals as finsets and fintypes. open finset int instance : locally_finite_order ℤ := -{ finset_Icc := λ a b, (Iio (b + 1 - a).to_nat).map ⟨λ n, n + a, λ _, by simp only - [imp_self, forall_const, add_left_inj, int.coe_nat_inj']⟩, +have hinj : ∀ a : ℤ, function.injective (λ n : ℕ, (n : ℤ) + a), +from λ a m n, by { rw [add_left_inj, coe_nat_inj'], exact id }, +{ finset_Icc := λ a b, (finset.range (b + 1 - a).to_nat).map + ⟨λ n, n + a, hinj a⟩, + finset_Ico := λ a b, (finset.range (b - a).to_nat).map + ⟨λ n, n + a, hinj a⟩, + finset_Ioc := λ a b, (finset.range (b - a).to_nat).map + ⟨λ n, n + (a + 1), hinj (a + 1)⟩, + finset_Ioo := λ a b, (finset.range (b - a - 1).to_nat).map + ⟨λ n, n + (a + 1), hinj (a + 1)⟩, finset_mem_Icc := λ a b x, begin - rw [mem_map, set.mem_Icc, function.embedding.coe_fn_mk], + simp only [int.lt_to_nat, exists_prop, mem_range, add_comm, function.embedding.coe_fn_mk, + mem_map], split, - { rintro ⟨n, hn, hx⟩, - rw [←hx, le_add_iff_nonneg_left], - rw [mem_Iio_iff, int.lt_to_nat, lt_sub_iff_add_lt, int.lt_add_one_iff] at hn, - exact ⟨int.coe_nat_nonneg _, hn⟩ }, - rintro h, - refine ⟨(x - a).to_nat, _, by rw [int.to_nat_sub_of_le h.1, int.sub_add_cancel]⟩, - rw mem_Iio_iff, - have hb := int.lt_add_one_of_le h.2, - exact (int.to_nat_lt_to_nat (sub_pos.2 (h.1.trans_lt hb))).2 (sub_lt_sub_right hb _), + { rintro ⟨a, h, rfl⟩, + rw [lt_sub_iff_add_lt, int.lt_add_one_iff, add_comm] at h, + refine ⟨int.le.intro rfl, h⟩ }, + { rintro ⟨ha, hb⟩, + use (x - a).to_nat, + rw to_nat_sub_of_le ha, + rw ←lt_add_one_iff at hb, + exact ⟨sub_lt_sub_right hb _, add_sub_cancel'_right _ _⟩ } + end, + finset_mem_Ico := λ a b x, begin + simp only [int.lt_to_nat, exists_prop, mem_range, add_comm, function.embedding.coe_fn_mk, + mem_map], + split, + { rintro ⟨a, h, rfl⟩, + exact ⟨int.le.intro rfl, lt_sub_iff_add_lt'.mp h⟩ }, + { rintro ⟨ha, hb⟩, + use (x - a).to_nat, + rw to_nat_sub_of_le ha, + exact ⟨sub_lt_sub_right hb _, add_sub_cancel'_right _ _⟩ } + end, + finset_mem_Ioc := λ a b x, begin + simp only [int.lt_to_nat, exists_prop, mem_range, add_comm, function.embedding.coe_fn_mk, + mem_map], + split, + { rintro ⟨a, h, rfl⟩, + refine ⟨int.le.intro rfl, _⟩, + rwa [←add_one_le_iff, le_sub_iff_add_le', add_comm _ (1 : ℤ), ←add_assoc] at h }, + { rintro ⟨ha, hb⟩, + use (x - (a + 1)).to_nat, + rw [to_nat_sub_of_le ha, ←add_one_le_iff, sub_add, add_sub_cancel], + exact ⟨sub_le_sub_right hb _, add_sub_cancel'_right _ _⟩ } + end, + finset_mem_Ioo := λ a b x, begin + simp only [int.lt_to_nat, exists_prop, mem_range, add_comm, function.embedding.coe_fn_mk, + mem_map], + split, + { rintro ⟨a, h, rfl⟩, + refine ⟨int.le.intro rfl, _⟩, + rwa [sub_sub, lt_sub_iff_add_lt'] at h }, + { rintro ⟨ha, hb⟩, + use (x - (a + 1)).to_nat, + rw [to_nat_sub_of_le ha, sub_sub], + exact ⟨sub_lt_sub_right hb _, add_sub_cancel'_right _ _⟩ } end } namespace int -variables (a b : ℤ) - -/-- `int_Ico a b` is the set of integers `b ≤ k < a`. -/ -def int_Ico (a b : ℤ) : finset ℤ := -(finset.range (b - a).to_nat).map - { to_fun := λ n, n + b, - inj' := λ n m h, by simpa using h } - -@[simp] lemma int_Ico.mem {n m l : ℤ} : l ∈ int_Ico n m ↔ n ≤ l ∧ l < m := -begin - dsimp [int_Ico], - simp only [int.lt_to_nat, exists_prop, mem_range, add_comm, function.embedding.coe_fn_mk, - mem_map], - split, - { rintro ⟨a, ⟨h, rfl⟩⟩, - exact ⟨int.le.intro rfl, lt_sub_iff_add_lt'.mp h⟩ }, - { rintro ⟨h₁, h₂⟩, - use (l - n).to_nat, - split; simp [h₁, h₂], } -end - -@[simp] lemma int_Ico.card (a b : ℤ) : (int_Ico a b).card = (b - a).to_nat := by simp [int_Ico] +variables (a b : ℤ). -@[simp] lemma card_finset_Icc : (Icc a b).card = (b + 1 - a).to_nat := sorry +@[simp] lemma card_finset_Icc : (Icc a b).card = (b + 1 - a).to_nat := +by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } -@[simp] lemma card_finset_Ico : (Ico a b).card = (b - a).to_nat := sorry +@[simp] lemma card_finset_Ico : (Ico a b).card = (b - a).to_nat := +by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } -@[simp] lemma card_finset_Ioc : (Ioc a b).card = (b - a).to_nat := sorry +@[simp] lemma card_finset_Ioc : (Ioc a b).card = (b - a).to_nat := +by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } -@[simp] lemma card_finset_Ioo : (Ioo a b).card = (b - a - 1).to_nat := sorry +@[simp] lemma card_finset_Ioo : (Ioo a b).card = (b - a - 1).to_nat := +by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = (b + 1 - a).to_nat := by rw [←card_finset_Icc, fintype.card_of_finset] diff --git a/src/data/nat/intervals.lean b/src/data/nat/intervals.lean index f6907f819de87..878bc016596cb 100644 --- a/src/data/nat/intervals.lean +++ b/src/data/nat/intervals.lean @@ -17,24 +17,64 @@ open finset nat instance : locally_finite_order ℕ := { finset_Icc := λ a b, (list.range' a (b + 1 - a)).to_finset, + finset_Ico := λ a b, (list.range' a (b - a)).to_finset, + finset_Ioc := λ a b, (list.range' (a + 1) (b - a)).to_finset, + finset_Ioo := λ a b, (list.range' (a + 1) (b - a - 1)).to_finset, finset_mem_Icc := λ a b x, begin rw [list.mem_to_finset, list.mem_range'], cases le_or_lt a b, { rw [nat.add_sub_cancel' (nat.lt_succ_of_le h).le, nat.lt_succ_iff] }, { rw [nat.sub_eq_zero_iff_le.2 h, add_zero], exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)) } -end } + end, + finset_mem_Ico := λ a b x, begin + rw [list.mem_to_finset, list.mem_range'], + cases le_or_lt a b, + { rw [nat.add_sub_cancel' h] }, + { rw [nat.sub_eq_zero_iff_le.2 h.le, add_zero], + exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2.le)) } + end, + finset_mem_Ioc := λ a b x, begin + rw [list.mem_to_finset, list.mem_range'], + cases le_or_lt a b, + { rw [←succ_sub_succ, nat.add_sub_cancel' (succ_le_succ h), nat.lt_succ_iff, nat.succ_le_iff] }, + { rw [nat.sub_eq_zero_iff_le.2 h.le, add_zero], + exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.le.trans hx.2)) } + end, + finset_mem_Ioo := λ a b x, begin + rw [list.mem_to_finset, list.mem_range', nat.sub_sub], + cases le_or_lt (a + 1) b, + { rw [nat.add_sub_cancel' h, nat.succ_le_iff] }, + { rw [nat.sub_eq_zero_iff_le.2 h.le, add_zero], + exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)) } + end } namespace nat variables (a b : ℕ) -@[simp] lemma card_finset_Icc : (Icc a b).card = b + 1 - a := sorry +@[simp] lemma card_finset_Icc : (Icc a b).card = b + 1 - a := +begin + change (list.to_finset _).card = _, + rw [list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'], +end -@[simp] lemma card_finset_Ico : (Ico a b).card = b - a := sorry +@[simp] lemma card_finset_Ico : (Ico a b).card = b - a := +begin + change (list.to_finset _).card = _, + rw [list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'], +end -@[simp] lemma card_finset_Ioc : (Ioc a b).card = b - a := sorry +@[simp] lemma card_finset_Ioc : (Ioc a b).card = b - a := +begin + change (list.to_finset _).card = _, + rw [list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'], +end -@[simp] lemma card_finset_Ioo : (Ioo a b).card = b - a - 1 := sorry +@[simp] lemma card_finset_Ioo : (Ioo a b).card = b - a - 1 := +begin + change (list.to_finset _).card = _, + rw [list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'], +end @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := by rw [←card_finset_Icc, fintype.card_of_finset] diff --git a/src/data/nat/multiplicity.lean b/src/data/nat/multiplicity.lean index 775d4fdd35484..bf47291d68b80 100644 --- a/src/data/nat/multiplicity.lean +++ b/src/data/nat/multiplicity.lean @@ -55,7 +55,7 @@ calc ... = ↑((finset.Ico 1 b).filter (λ i, m ^ i ∣ n)).card : congr_arg coe $ congr_arg card $ finset.ext $ λ i, begin - rw [mem_filter, Ico.mem, Ico.mem, lt_succ_iff, ←@enat.coe_le_coe i, enat.coe_get, + rw [mem_filter, mem_Ico, mem_Ico, lt_succ_iff, ←@enat.coe_le_coe i, enat.coe_get, ←pow_dvd_iff_le_multiplicity, and.right_comm], refine (and_iff_left_of_imp (λ h, _)).symm, cases m, diff --git a/src/data/nat/totient.lean b/src/data/nat/totient.lean index d7a719be77415..3fdfe8b114798 100644 --- a/src/data/nat/totient.lean +++ b/src/data/nat/totient.lean @@ -187,10 +187,10 @@ begin have hempty : finset.filter p.coprime {0} = ∅, { simp only [finset.filter_singleton, nat.coprime_zero_right, hp.ne', if_false] }, rw [totient_eq_card_coprime, hsplit, finset.filter_union, hempty, finset.empty_union, - ←finset.Ico.card 1 p] at h, + ←nat.card_finset_Ico 1 p] at h, refine p.prime_of_coprime hp (λ n hn hnz, _), apply finset.filter_card_eq h n, - refine finset.Ico.mem.mpr ⟨_, hn⟩, + refine finset.mem_Ico.mpr ⟨_, hn⟩, rwa [succ_le_iff, pos_iff_ne_zero], end diff --git a/src/data/pnat/intervals.lean b/src/data/pnat/intervals.lean index 6f6398b4f9c62..933196019aaa6 100644 --- a/src/data/pnat/intervals.lean +++ b/src/data/pnat/intervals.lean @@ -3,8 +3,7 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Yaël Dillies -/ -import data.pnat.basic -import order.locally_finite +import data.nat.intervals /-! # Intervals of positive naturals @@ -15,36 +14,38 @@ intervals as finsets and fintypes. open finset pnat -instance : locally_finite_order ℕ+ := -{ finset_Icc := λ a b, (list.range' a (b + 1 - a)).to_finset, - finset_mem_Icc := λ a b x, begin - rw [list.mem_to_finset, list.mem_range'], - cases le_or_lt a b, - { rw [nat.add_sub_cancel' (nat.lt_succ_of_le h).le, nat.lt_succ_iff] }, - { rw [nat.sub_eq_zero_iff_le.2 h, add_zero], - exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)) } -end } +instance : locally_finite_order ℕ+ := subtype.locally_finite_order _ namespace pnat variables (a b : ℕ+) -/-- `Ico a b` is the set of positive natural numbers `b ≤ k < a`. -/ -def pnat_Ico (a b : ℕ+) : finset ℕ+ := -(finset.Ico a b).attach.map - { to_fun := λ n, ⟨(n : ℕ), lt_of_lt_of_le a.2 (finset.Ico.mem.1 n.2).1⟩, - -- why can't we do this directly? - inj' := λ n m h, subtype.eq (by { replace h := congr_arg subtype.val h, exact h }) } - -@[simp] lemma pnat_Ico.mem : ∀ {n m l : ℕ+}, l ∈ pnat_Ico n m ↔ n ≤ l ∧ l < m := -by { rintro ⟨n, hn⟩ ⟨m, hm⟩ ⟨l, hl⟩, simp [pnat_Ico] } - -@[simp] lemma card_finset_Icc : (Icc a b).card = b + 1 - a := sorry - -@[simp] lemma card_finset_Ico : (Ico a b).card = b - a := sorry - -@[simp] lemma card_finset_Ioc : (Ioc a b).card = b - a := sorry - -@[simp] lemma card_finset_Ioo : (Ioo a b).card = b - a - 1 := sorry +@[simp] lemma card_finset_Icc : (Icc a b).card = b + 1 - a := +begin + change (finset.subtype _ _).card = _, + rw [finset.card_subtype, finset.filter_true_of_mem, nat.card_finset_Icc], refl, + exact λ x hx, a.2.trans_le (mem_Icc.1 hx).1, +end + +@[simp] lemma card_finset_Ico : (Ico a b).card = b - a := +begin + change (finset.subtype _ _).card = _, + rw [finset.card_subtype, finset.filter_true_of_mem, nat.card_finset_Ico], refl, + exact λ x hx, a.2.trans_le (mem_Ico.1 hx).1, +end + +@[simp] lemma card_finset_Ioc : (Ioc a b).card = b - a := +begin + change (finset.subtype _ _).card = _, + rw [finset.card_subtype, finset.filter_true_of_mem, nat.card_finset_Ioc], refl, + exact λ x hx, a.2.trans (mem_Ioc.1 hx).1, +end + +@[simp] lemma card_finset_Ioo : (Ioo a b).card = b - a - 1 := +begin + change (finset.subtype _ _).card = _, + rw [finset.card_subtype, finset.filter_true_of_mem, nat.card_finset_Ioo], refl, + exact λ x hx, a.2.trans (mem_Ioo.1 hx).1, +end @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := by rw [←card_finset_Icc, fintype.card_of_finset] diff --git a/src/data/polynomial/inductions.lean b/src/data/polynomial/inductions.lean index 2c7efa041b9a7..c01a3ce77a882 100644 --- a/src/data/polynomial/inductions.lean +++ b/src/data/polynomial/inductions.lean @@ -32,7 +32,7 @@ def div_X (p : polynomial R) : polynomial R := @[simp] lemma coeff_div_X : (div_X p).coeff n = p.coeff (n+1) := begin simp only [div_X, coeff_monomial, true_and, finset_sum_coeff, not_lt, - Ico.mem, zero_le, finset.sum_ite_eq', ite_eq_left_iff], + mem_Ico, zero_le, finset.sum_ite_eq', ite_eq_left_iff], intro h, rw coeff_eq_zero_of_nat_degree_lt (nat.lt_succ_of_le h) end diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index c62ed005b00df..96dc0f49eef04 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -3,9 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import data.finset.lattice import data.finset.preimage -import data.finset.sort import data.set.finite import data.set.intervals.basic @@ -46,17 +44,24 @@ When it's also an `order_bot`, ## Instances We provide a `locally_finite_order` instance for -* ℕ, see `nat.locally_finite_order` -* ℤ, see `int.locally_finite_order` * a subtype of a locally finite order * any fintype (but it is noncomputable), see `fintype.locally_finite_order` +Instances for concrete types are proved in their respective files: +* `data.nat.intervals` for `ℕ` +* `data.int.intervals` for `ℤ` +* `data.pnat.intervals` for `ℕ+` +Along, you will find lemmas about the cardinality of those finite intervals. + ## TODO -TODO: All the API in `set.intervals.basic` could be copied here. But this is unfortunately labor- -and maintenance-intensive. +Provide the `locally_finite_order` instance for `lex α β`.z + +From `linear_order α`, `no_top_order α`, `locally_finite_order α`, we can also define an +order isomorphism `α ≃ ℕ` or `α ≃ ℤ`, depending on whether we have `order_bot α` or +`no_bot_order α` and `nonempty α`. When `order_bot α`, we can match `a : α` to `(Iio a).card`. -TODO: Once we have the `succ_order` (any non-top element has a least greater element) typeclass, we +Once we have the `succ_order` typeclass (any non-top element has a least greater element), we can provide `succ_order α` from `linear_order α` and `locally_finite_order α` using ```lean @@ -75,37 +80,18 @@ begin -- very non golfed exact (finset.min'_le _ _ (finset.mem_Ioc_iff.2 ⟨hx, le_rfl⟩)).trans hy, end ``` - Note that the converse is not true. Consider `{-2^z | z : ℤ} ∪ {2^z | z : ℤ}`. Any element has a -successor, so it is a `succ_order`, but it's not locally finite as `Icc (-1) 1` is infinite. - -TODO: From `linear_order α`, `no_top_order α`, `locally_finite_order α`, we can also define an -order isomorphism `α ≃ ℕ` or `α ≃ ℤ`, depending on whether we have `order_bot α` or -`no_bot_order α` and `nonempty α`. When `order_bot α`, we can match `a : α` to `(Iio a).card`. +successor (and actually a predecessor as well), so it is a `succ_order`, but it's not locally finite +as `Icc (-1) 1` is infinite. -/ lemma and_and_and_comm (a b c d : Prop) : (a ∧ b) ∧ c ∧ d ↔ (a ∧ c) ∧ b ∧ d := by rw [←and_assoc, @and.right_comm a, and_assoc] -section tactic -open tactic - -meta def finset_Ico_laws_tac : tactic unit := -whnf_target >> intros >> to_expr``(begin - unfold finset_Ico, - rw [finset.mem_filter, finset_mem_Icc, and_assoc, lt_iff_le_not_le] -end) >>= exact - -meta def finset_Ioc_laws_tac : tactic unit := -whnf_target >> intros >> to_expr``(by { unfold locally_finite_order.finset_Ioc, - rw [finset.mem_filter, finset_mem_Icc, and.right_comm, lt_iff_le_not_le] }) >>= exact - -meta def finset_Ioo_laws_tac : tactic unit := -whnf_target >> intros >> to_expr``(by { unfold locally_finite_order.finset_Ioo, - rw [finset.mem_filter, finset_mem_Icc, and_and_and_comm, lt_iff_le_not_le, lt_iff_le_not_le] }) - >>= exact - -end tactic +@[simp] protected lemma list.nodup.erase_dup {α : Type*} [decidable_eq α] {l : list α} + (h : l.nodup) : + l.erase_dup = l := +list.erase_dup_eq_self.2 h open finset @@ -115,9 +101,9 @@ class locally_finite_order (α : Type*) [preorder α] [decidable_rel ((≤) : α (finset_Ioc : α → α → finset α := λ a b, (finset_Icc a b).filter (λ x, ¬x ≤ a)) (finset_Ioo : α → α → finset α := λ a b, (finset_Icc a b).filter (λ x, ¬x ≤ a ∧ ¬b ≤ x)) (finset_mem_Icc : ∀ a b x : α, x ∈ finset_Icc a b ↔ a ≤ x ∧ x ≤ b) -(finset_mem_Ico : ∀ a b x : α, x ∈ finset_Ico a b ↔ a ≤ x ∧ x < b . finset_Ico_laws_tac) -(finset_mem_Ioc : ∀ a b x : α, x ∈ finset_Ioc a b ↔ a < x ∧ x ≤ b . finset_Ioc_laws_tac) -(finset_mem_Ioo : ∀ a b x : α, x ∈ finset_Ioo a b ↔ a < x ∧ x < b . finset_Ioo_laws_tac) +(finset_mem_Ico : ∀ a b x : α, x ∈ finset_Ico a b ↔ a ≤ x ∧ x < b) +(finset_mem_Ioc : ∀ a b x : α, x ∈ finset_Ioc a b ↔ a < x ∧ x ≤ b) +(finset_mem_Ioo : ∀ a b x : α, x ∈ finset_Ioo a b ↔ a < x ∧ x < b) variables {α β : Type*} @@ -381,53 +367,62 @@ end set /-! ### Instances -/ -section +open finset + +section preorder variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] noncomputable def locally_finite_order_of_finite_Icc (h : ∀ a b : α, (set.Icc a b).finite) : locally_finite_order α := { finset_Icc := λ a b, (h a b).to_finset, + finset_Ico := λ a b, ((h a b).subset set.Ico_subset_Icc_self).to_finset, + finset_Ioc := λ a b, ((h a b).subset set.Ioc_subset_Icc_self).to_finset, + finset_Ioo := λ a b, ((h a b).subset set.Ioo_subset_Icc_self).to_finset, finset_mem_Icc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Icc], - finset_mem_Ico := λ a b x, begin - rw [finset.mem_filter, finset_mem_Icc, and_assoc, lt_iff_le_not_le], - end } + finset_mem_Ico := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ico], + finset_mem_Ioc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioc], + finset_mem_Ioo := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioo] } -@[priority 900] -noncomputable instance fintype.locally_finite_order [fintype α] : +noncomputable def fintype.to_locally_finite_order [fintype α] : locally_finite_order α := { finset_Icc := λ a b, (set.finite.of_fintype (set.Icc a b)).to_finset, - finset_mem_Icc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Icc] } + finset_Ico := λ a b, (set.finite.of_fintype (set.Ico a b)).to_finset, + finset_Ioc := λ a b, (set.finite.of_fintype (set.Ioc a b)).to_finset, + finset_Ioo := λ a b, (set.finite.of_fintype (set.Ioo a b)).to_finset, + finset_mem_Icc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Icc], + finset_mem_Ico := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ico], + finset_mem_Ioc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioc], + finset_mem_Ioo := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioo] } --- Should this be called `order_embedding.locally_finite_order`? +instance [locally_finite_order α] (p : α → Prop) [decidable_pred p] : + locally_finite_order (subtype p) := +{ finset_Icc := λ a b, (Icc a.val b.val).subtype p, + finset_Ico := λ a b, (Ico a.val b.val).subtype p, + finset_Ioc := λ a b, (Ioc a.val b.val).subtype p, + finset_Ioo := λ a b, (Ioo a.val b.val).subtype p, + finset_mem_Icc := λ a b x, by simp_rw [finset.mem_subtype, mem_Icc, subtype.val_eq_coe, + subtype.coe_le_coe], + finset_mem_Ico := λ a b x, by simp_rw [finset.mem_subtype, mem_Ico, subtype.val_eq_coe, + subtype.coe_le_coe, subtype.coe_lt_coe], + finset_mem_Ioc := λ a b x, by simp_rw [finset.mem_subtype, mem_Ioc, subtype.val_eq_coe, + subtype.coe_le_coe, subtype.coe_lt_coe], + finset_mem_Ioo := λ a b x, by simp_rw [finset.mem_subtype, mem_Ioo, subtype.val_eq_coe, + subtype.coe_lt_coe] } + +variables [preorder β] [decidable_rel ((≤) : β → β → Prop)] [locally_finite_order β] + +-- Should this be called `locally_finite_order.lift`? /-- Given an order embedding `α ↪o β`, pulls back the `locally_finite_order` on `β` to `α`. -/ -noncomputable def locally_finite_order.lift {β : Type*} [partial_order β] [locally_finite_order β] - [decidable_eq β] [partial_order α] [decidable_eq α] (f : α ↪o β) : +noncomputable def order_embedding.locally_finite_order (f : α ↪o β) : locally_finite_order α := { finset_Icc := λ a b, (Icc (f a) (f b)).preimage f (f.to_embedding.injective.inj_on _), - finset_mem_Icc := λ a b x, by rw [mem_preimage, mem_Icc, f.le_iff_le, f.le_iff_le] } + finset_Ico := λ a b, (Ico (f a) (f b)).preimage f (f.to_embedding.injective.inj_on _), + finset_Ioc := λ a b, (Ioc (f a) (f b)).preimage f (f.to_embedding.injective.inj_on _), + finset_Ioo := λ a b, (Ioo (f a) (f b)).preimage f (f.to_embedding.injective.inj_on _), + finset_mem_Icc := λ a b x, by rw [mem_preimage, mem_Icc, f.le_iff_le, f.le_iff_le], + finset_mem_Ico := λ a b x, by rw [mem_preimage, mem_Ico, f.le_iff_le, f.lt_iff_lt], + finset_mem_Ioc := λ a b x, by rw [mem_preimage, mem_Ioc, f.lt_iff_lt, f.le_iff_le], + finset_mem_Ioo := λ a b x, by rw [mem_preimage, mem_Ioo, f.lt_iff_lt, f.lt_iff_lt] } -instance subtype.locally_finite_order [locally_finite_order α] {p : α → Prop} [decidable_pred p] : - locally_finite_order (subtype p) := -begin - haveI : preorder (subtype p) := by apply_instance, - haveI : decidable_rel ((≤) : subtype p → subtype p → Prop) := by apply_instance, - exact - { finset_Icc := λ a b, finset.subtype p (Icc a.val b.val), - finset_mem_Icc := λ a b x, begin - rw [finset.subtype, mem_map, set.mem_image, set.mem_Icc], - dsimp, - split, - { rintro ⟨⟨y, hy⟩, -, h⟩, - rw [←h, ←subtype.coe_le_coe, ←subtype.coe_le_coe], - rw [mem_filter, mem_Icc_iff] at hy, - exact hy.1 }, - rintro hx, - refine ⟨⟨x, _⟩, mem_coe.2 (mem_attach _ _), _⟩, - { rw [mem_filter, mem_Icc_iff, subtype.coe_le_coe, subtype.coe_le_coe], - exact ⟨hx, x.2⟩ }, - simp only [subtype.coe_eta, subtype.coe_mk], - end }, -end - -end +end preorder diff --git a/src/tactic/interval_cases.lean b/src/tactic/interval_cases.lean index 85349e969e337..5da1d9006bc7c 100644 --- a/src/tactic/interval_cases.lean +++ b/src/tactic/interval_cases.lean @@ -26,8 +26,8 @@ The hypotheses should be in the form `hl : a ≤ n` and `hu : n < b`, in which case `interval_cases` calls `fin_cases` on the resulting fact `n ∈ set.Ico a b`. -/ +import data.nat.intervals import tactic.fin_cases -import data.fintype.intervals open set diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index b6f621a18789c..b9bf737740199 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -6,7 +6,7 @@ Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébas import topology.metric_space.emetric_space import topology.algebra.ordered.basic -import data.fintype.intervals +import data.nat.intervals /-! # Metric spaces From 80e9a4d4f05338d8b78a6153fd580f681e38bab9 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Tue, 14 Sep 2021 22:54:33 +0200 Subject: [PATCH 08/41] elaboration witchcraft and prod instance --- src/data/list/erase_dup.lean | 3 ++ src/logic/basic.lean | 6 +++ src/order/locally_finite.lean | 90 ++++++++++++++++++++++++++++------- 3 files changed, 82 insertions(+), 17 deletions(-) diff --git a/src/data/list/erase_dup.lean b/src/data/list/erase_dup.lean index 38a4a9619905c..fae6ec978d71a 100644 --- a/src/data/list/erase_dup.lean +++ b/src/data/list/erase_dup.lean @@ -56,6 +56,9 @@ theorem nodup_erase_dup : ∀ l : list α, nodup (erase_dup l) := pairwise_pw_fi theorem erase_dup_eq_self {l : list α} : erase_dup l = l ↔ nodup l := pw_filter_eq_self +@[simp] protected lemma list.nodup.erase_dup {l : list α} (h : l.nodup) : l.erase_dup = l := +list.erase_dup_eq_self.2 h + @[simp] theorem erase_dup_idempotent {l : list α} : erase_dup (erase_dup l) = erase_dup l := pw_filter_idempotent diff --git a/src/logic/basic.lean b/src/logic/basic.lean index 7e87d7748c328..cf6022a90105c 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -161,6 +161,9 @@ lemma ne_comm {α} {a b : α} : a ≠ b ↔ b ≠ a := ⟨ne.symm, ne.symm⟩ (∀ {c}, a = c ↔ b = c) ↔ (a = b) := ⟨λ h, by rw h, λ h a, by rw h⟩ +/-- Annotation identity function to force Lean to see `a` in the type of the goal. -/ +abbreviation id_annotate {α} (a : α) {β} (b : β) := b + /-- Wrapper for adding elementary propositions to the type class systems. Warning: this can easily be abused. See the rest of this docstring for details. @@ -383,6 +386,9 @@ and.imp id h lemma and.right_comm : (a ∧ b) ∧ c ↔ (a ∧ c) ∧ b := by simp only [and.left_comm, and.comm] +lemma and_and_and_comm (a b c d : Prop) : (a ∧ b) ∧ c ∧ d ↔ (a ∧ c) ∧ b ∧ d := +by rw [←and_assoc, @and.right_comm a, and_assoc] + lemma and.rotate : a ∧ b ∧ c ↔ b ∧ c ∧ a := by simp only [and.left_comm, and.comm] diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 96dc0f49eef04..54db2f582df34 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -55,7 +55,7 @@ Along, you will find lemmas about the cardinality of those finite intervals. ## TODO -Provide the `locally_finite_order` instance for `lex α β`.z +Provide the `locally_finite_order` instance for `lex α β`. From `linear_order α`, `no_top_order α`, `locally_finite_order α`, we can also define an order isomorphism `α ≃ ℕ` or `α ≃ ℤ`, depending on whether we have `order_bot α` or @@ -85,13 +85,66 @@ successor (and actually a predecessor as well), so it is a `succ_order`, but it' as `Icc (-1) 1` is infinite. -/ -lemma and_and_and_comm (a b c d : Prop) : (a ∧ b) ∧ c ∧ d ↔ (a ∧ c) ∧ b ∧ d := -by rw [←and_assoc, @and.right_comm a, and_assoc] +/-! +### Arguments autofilling tactics + +`locally_finite_order` has fields for all of `Icc`, `Ico`, `Ioc`, `Ioo` to let control over +definitional equality. But all fields can be recovered from `Icc` (or any other actually, but we had +to make a choice here). Hence we provide default values for `locally_finite_order.finset_Ico`, +`locally_finite_order.finset_Ioc`, `locally_finite_order.finset_Ioo` and three corresponding tactics +to fill in the three corresponding proofs. + +This is analogous to the default definition of `(<)` in terms of `(≤)`. There, the default proof of +`lt_iff_le_not_le` is provided by `order_laws_tac`, which just tries to call `refl`. Here, the +situation is more complicated, as we need to call an hypothesis (namely +`locally_finite_order.mem_finset_Icc`) that doesn't appear a priori in the type of the proof. Whence +our use of `id_annotate` to force it to appear there. +-/ -@[simp] protected lemma list.nodup.erase_dup {α : Type*} [decidable_eq α] {l : list α} - (h : l.nodup) : - l.erase_dup = l := -list.erase_dup_eq_self.2 h +section tactic +open tactic + +variables {α : Type*} [preorder α] [decidable_rel ((≤) : α → α → Prop)] + {finset_Icc : α → α → finset α} (finset_mem_Icc : ∀ a b x : α, x ∈ finset_Icc a b ↔ a ≤ x ∧ x ≤ b) + +/-- Auxiliary lemma to provide `locally_finite_order.mem_finset_Ico` when +`locally_finite_order.finset_Ico` is the default value. -/ +lemma locally_finite_order.mem_Ico_of_mem_Icc : + id_annotate finset_mem_Icc + (∀ a b x : α, x ∈ (finset_Icc a b).filter (λ x, ¬b ≤ x) ↔ a ≤ x ∧ x < b) := +λ a b x, by rw [finset.mem_filter, finset_mem_Icc, and_assoc, lt_iff_le_not_le] + +/-- Auxiliary lemma to provide `locally_finite_order.mem_finset_Ioc` when +`locally_finite_order.finset_Ioc` is the default value. -/ +lemma locally_finite_order.mem_Ioc_of_mem_Icc : + id_annotate finset_mem_Icc + (∀ a b x : α, x ∈ (finset_Icc a b).filter (λ x, ¬x ≤ a) ↔ a < x ∧ x ≤ b) := +λ a b x, by rw [finset.mem_filter, finset_mem_Icc, and.right_comm, lt_iff_le_not_le] + +/-- Auxiliary lemma to provide `locally_finite_order.mem_finset_Ioo` when +`locally_finite_order.finset_Ico` is the default value. -/ +lemma locally_finite_order.mem_Ioo_of_mem_Icc : + id_annotate finset_mem_Icc + (∀ a b x : α, x ∈ (finset_Icc a b).filter (λ x, ¬x ≤ a ∧ ¬b ≤ x) ↔ a < x ∧ x < b) := +λ a b x, by rw [finset.mem_filter, finset_mem_Icc, and_and_and_comm, lt_iff_le_not_le, + lt_iff_le_not_le] + +/-- Here to make `locally_finite_order.finset_Ico` an optional field. Tries to prove +`locally_finite_order.mem_finset_Ico` -/ +meta def finset_Ico_laws_tac : tactic unit := +`[exact locally_finite_order.mem_Ico_of_mem_Icc _] + +/-- Here to make `locally_finite_order.finset_Ioc` an optional field. Tries to prove +`locally_finite_order.mem_finset_Ioc` -/ +meta def finset_Ioc_laws_tac : tactic unit := +`[exact locally_finite_order.mem_Ioc_of_mem_Icc _] + +/-- Here to make `locally_finite_order.finset_Ioo` an optional field. Tries to prove +`locally_finite_order.mem_finset_Ioo` -/ +meta def finset_Ioo_laws_tac : tactic unit := +`[exact locally_finite_order.mem_Ioo_of_mem_Icc _] + +end tactic open finset @@ -101,9 +154,12 @@ class locally_finite_order (α : Type*) [preorder α] [decidable_rel ((≤) : α (finset_Ioc : α → α → finset α := λ a b, (finset_Icc a b).filter (λ x, ¬x ≤ a)) (finset_Ioo : α → α → finset α := λ a b, (finset_Icc a b).filter (λ x, ¬x ≤ a ∧ ¬b ≤ x)) (finset_mem_Icc : ∀ a b x : α, x ∈ finset_Icc a b ↔ a ≤ x ∧ x ≤ b) -(finset_mem_Ico : ∀ a b x : α, x ∈ finset_Ico a b ↔ a ≤ x ∧ x < b) -(finset_mem_Ioc : ∀ a b x : α, x ∈ finset_Ioc a b ↔ a < x ∧ x ≤ b) -(finset_mem_Ioo : ∀ a b x : α, x ∈ finset_Ioo a b ↔ a < x ∧ x < b) +(finset_mem_Ico : id_annotate finset_mem_Icc (∀ a b x : α, x ∈ finset_Ico a b ↔ a ≤ x ∧ x < b) + . finset_Ico_laws_tac) +(finset_mem_Ioc : id_annotate finset_mem_Icc (∀ a b x : α, x ∈ finset_Ioc a b ↔ a < x ∧ x ≤ b) + . finset_Ioc_laws_tac) +(finset_mem_Ioo : id_annotate finset_mem_Icc (∀ a b x : α, x ∈ finset_Ioo a b ↔ a < x ∧ x < b) + . finset_Ioo_laws_tac) variables {α β : Type*} @@ -376,13 +432,7 @@ noncomputable def locally_finite_order_of_finite_Icc (h : ∀ a b : α, (set.Icc a b).finite) : locally_finite_order α := { finset_Icc := λ a b, (h a b).to_finset, - finset_Ico := λ a b, ((h a b).subset set.Ico_subset_Icc_self).to_finset, - finset_Ioc := λ a b, ((h a b).subset set.Ioc_subset_Icc_self).to_finset, - finset_Ioo := λ a b, ((h a b).subset set.Ioo_subset_Icc_self).to_finset, - finset_mem_Icc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Icc], - finset_mem_Ico := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ico], - finset_mem_Ioc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioc], - finset_mem_Ioo := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioo] } + finset_mem_Icc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Icc] } noncomputable def fintype.to_locally_finite_order [fintype α] : locally_finite_order α := @@ -425,4 +475,10 @@ noncomputable def order_embedding.locally_finite_order (f : α ↪o β) : finset_mem_Ioc := λ a b x, by rw [mem_preimage, mem_Ioc, f.lt_iff_lt, f.le_iff_le], finset_mem_Ioo := λ a b x, by rw [mem_preimage, mem_Ioo, f.lt_iff_lt, f.lt_iff_lt] } +variables [locally_finite_order α] + +instance : locally_finite_order (α × β) := +{ finset_Icc := λ a b, (Icc a.fst b.fst).product (Icc a.snd b.snd), + finset_mem_Icc := λ a b x, by { rw [mem_product, mem_Icc, mem_Icc, and_and_and_comm], refl } } + end preorder From 372b00cc5fb79d5b7bf2fdffad0e397231c15384 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Wed, 15 Sep 2021 16:14:06 +0200 Subject: [PATCH 09/41] fix list.nodup.erase_dup --- src/data/list/erase_dup.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/list/erase_dup.lean b/src/data/list/erase_dup.lean index fae6ec978d71a..6cf545f624532 100644 --- a/src/data/list/erase_dup.lean +++ b/src/data/list/erase_dup.lean @@ -56,7 +56,7 @@ theorem nodup_erase_dup : ∀ l : list α, nodup (erase_dup l) := pairwise_pw_fi theorem erase_dup_eq_self {l : list α} : erase_dup l = l ↔ nodup l := pw_filter_eq_self -@[simp] protected lemma list.nodup.erase_dup {l : list α} (h : l.nodup) : l.erase_dup = l := +@[simp] protected lemma nodup.erase_dup {l : list α} (h : l.nodup) : l.erase_dup = l := list.erase_dup_eq_self.2 h @[simp] theorem erase_dup_idempotent {l : list α} : erase_dup (erase_dup l) = erase_dup l := From eac6dd11046d8e53727b11a5515d0c66ebe1a1fb Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Thu, 16 Sep 2021 11:24:01 +0200 Subject: [PATCH 10/41] adapt the easy part of data.finset.intervals --- src/algebra/big_operators/intervals.lean | 12 +- src/algebra/geom_sum.lean | 4 +- src/analysis/analytic/inverse.lean | 4 +- src/data/finset/intervals.lean | 429 ++++++++++-------- src/data/finset/mwe.lean | 0 src/data/nat/intervals.lean | 78 +++- src/data/nat/totient.lean | 2 +- .../decomposition/unsigned_hahn.lean | 4 +- src/number_theory/divisors.lean | 2 +- src/order/locally_finite.lean | 6 + src/topology/metric_space/basic.lean | 2 +- src/topology/metric_space/emetric_space.lean | 2 +- 12 files changed, 329 insertions(+), 216 deletions(-) create mode 100644 src/data/finset/mwe.lean diff --git a/src/algebra/big_operators/intervals.lean b/src/algebra/big_operators/intervals.lean index 79b809b4f466f..f272fbc3a6bad 100644 --- a/src/algebra/big_operators/intervals.lean +++ b/src/algebra/big_operators/intervals.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl -/ import algebra.big_operators.basic -import data.finset.intervals +import data.nat.intervals import tactic.linarith /-! @@ -34,7 +34,7 @@ lemma prod_Ico_add (f : ℕ → β) (m n k : ℕ) : lemma sum_Ico_succ_top {δ : Type*} [add_comm_monoid δ] {a b : ℕ} (hab : a ≤ b) (f : ℕ → δ) : (∑ k in Ico a (b + 1), f k) = (∑ k in Ico a b, f k) + f b := -by rw [Ico.succ_top hab, sum_insert right_not_mem_Ico, add_comm] +by rw [Ico_insert_right hab, sum_insert right_not_mem_Ico, add_comm] @[to_additive] lemma prod_Ico_succ_top {a b : ℕ} (hab : a ≤ b) (f : ℕ → β) : @@ -92,7 +92,7 @@ begin by_cases h : m ≤ n, { rw [← Ico.zero_bot, prod_Ico_add, zero_add, nat.sub_add_cancel h] }, { replace h : n ≤ m := le_of_not_ge h, - rw [Ico.eq_empty_of_le h, nat.sub_eq_zero_of_le h, range_zero, prod_empty, prod_empty] } + rw [Ico_eq_empty_of_le h, nat.sub_eq_zero_of_le h, range_zero, prod_empty, prod_empty] } end lemma prod_Ico_reflect (f : ℕ → β) (k : ℕ) {m n : ℕ} (h : m ≤ n + 1) : @@ -107,7 +107,7 @@ begin simp only [mem_Ico], rintros i ⟨ki, im⟩ j ⟨kj, jm⟩ Hij, rw [← nat.sub_sub_self (this _ im), Hij, nat.sub_sub_self (this _ jm)] }, - { simp [Ico.eq_empty_of_le, nat.sub_le_sub_left, hkm] } + { simp [Ico_eq_empty_of_le, nat.sub_le_sub_left, hkm] } end lemma sum_Ico_reflect {δ : Type*} [add_comm_monoid δ] (f : ℕ → δ) (k : ℕ) {m n : ℕ} @@ -120,8 +120,8 @@ lemma prod_range_reflect (f : ℕ → β) (n : ℕ) : begin cases n, { simp }, - { simp only [range_eq_Ico, nat.succ_sub_succ_eq_sub, nat.sub_zero], - rw [prod_Ico_reflect _ _ (le_refl _)], + { simp only [←nat.Ico_zero_eq_range, nat.succ_sub_succ_eq_sub, nat.sub_zero], + rw prod_Ico_reflect _ _ le_rfl, simp } end diff --git a/src/algebra/geom_sum.lean b/src/algebra/geom_sum.lean index 8f7f93d90cac9..cbc0a1070dbb4 100644 --- a/src/algebra/geom_sum.lean +++ b/src/algebra/geom_sum.lean @@ -353,14 +353,14 @@ lemma nat.geom_sum_Ico_le {b : ℕ} (hb : 2 ≤ b) (a n : ℕ) : ∑ i in Ico 1 n, a/b^i ≤ a/(b - 1) := begin cases n, - { rw [Ico.eq_empty_of_le zero_le_one, sum_empty], + { rw [Ico_eq_empty_of_le zero_le_one, sum_empty], exact nat.zero_le _ }, rw ←add_le_add_iff_left a, calc a + ∑ (i : ℕ) in Ico 1 n.succ, a/b^i = a/b^0 + ∑ (i : ℕ) in Ico 1 n.succ, a/b^i : by rw [pow_zero, nat.div_one] ... = ∑ i in range n.succ, a/b^i : begin - rw [range_eq_Ico, ←finset.Ico.insert_succ_bot (nat.succ_pos _), sum_insert], + rw [range_eq_Ico, ←nat.Ico_insert_succ_left (nat.succ_pos _), sum_insert], exact λ h, zero_lt_one.not_le (mem_Ico.1 h).1, end ... ≤ a * b/(b - 1) : nat.geom_sum_le hb a _ diff --git a/src/analysis/analytic/inverse.lean b/src/analysis/analytic/inverse.lean index e696c34b5c6ca..c89e88a29aded 100644 --- a/src/analysis/analytic/inverse.lean +++ b/src/analysis/analytic/inverse.lean @@ -417,7 +417,7 @@ let I := ∥(i.symm : F →L[𝕜] E)∥ in calc ∑ k in Ico 1 (n + 1), a ^ k * ∥p.right_inv i k∥ = a * I + ∑ k in Ico 2 (n + 1), a ^ k * ∥p.right_inv i k∥ : by simp only [linear_isometry_equiv.norm_map, pow_one, right_inv_coeff_one, - Ico.succ_singleton, sum_singleton, ← sum_Ico_consecutive _ one_le_two hn] + Ico_succ_singleton, sum_singleton, ← sum_Ico_consecutive _ one_le_two hn] ... = a * I + ∑ k in Ico 2 (n + 1), a ^ k * ∥(i.symm : F →L[𝕜] E).comp_continuous_multilinear_map (∑ c in ({c | 1 < composition.length c}.to_finset : finset (composition k)), @@ -487,7 +487,7 @@ begin have IRec : ∀ n, 1 ≤ n → S n ≤ (I + 1) * a, { apply nat.le_induction, { simp only [S], - rw [Ico.eq_empty_of_le (le_refl 1), sum_empty], + rw [Ico_eq_empty_of_le (le_refl 1), sum_empty], exact mul_nonneg (add_nonneg (norm_nonneg _) zero_le_one) apos.le }, { assume n one_le_n hn, have In : 2 ≤ n + 1, by linarith, diff --git a/src/data/finset/intervals.lean b/src/data/finset/intervals.lean index c5fef0ed48143..ee9ae40ea3f78 100644 --- a/src/data/finset/intervals.lean +++ b/src/data/finset/intervals.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison +Authors: Scott Morrison, Yaël Dillies -/ import algebra.ordered_sub import data.finset.basic @@ -12,234 +12,307 @@ import order.locally_finite # Intervals as finsets For now this only covers `Ico a b`, the "closed-open" interval containing `[a, ..., b - 1]`. + +## TODO + +This file was originally only about `finset.Ico a b` where `a n : ℕ`. No care has yet been taken to +generalize these lemmas properly and many lemmas about `Icc`, `Ioc`, `Ioo` are missing. -/ -open finset nat +-- TODO: There is a diamond between `nat.decidable_eq` and `decidable_eq_of_decidable_le` +-- and another one between `linear_order.decidable_lt` and `decidable_lt_of_decidable_le` +local attribute [-instance] decidable_eq_of_decidable_le decidable_lt_of_decidable_le -section ordered_cancel_add_comm_monoid -variables {α : Type*} [ordered_cancel_add_comm_monoid α] [has_exists_add_of_le α] [has_sub α] - [has_ordered_sub α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] +lemma finset.coe_eq_singleton {α : Type*} {s : finset α} {a : α} : + (s : set α) = {a} ↔ s = {a} := +by rw [←finset.coe_singleton, finset.coe_inj] -lemma image_add_const_Icc (a b c : α) : (Icc a b).image ((+) c) = Icc (a + c) (b + c) := -begin - ext x, - rw [mem_image, mem_Icc], - split, - { rintro ⟨y, hy, hx⟩, - rw mem_Icc at hy, - rw [←hx, add_comm c], - exact ⟨add_le_add_right hy.1 c, add_le_add_right hy.2 c⟩ }, - intro hx, - obtain ⟨y, hy⟩ := exists_add_of_le hx.1, - rw [hy, add_right_comm] at hx, - rw [eq_comm, add_right_comm, add_comm] at hy, - exact ⟨a + y, mem_Icc.2 ⟨le_of_add_le_add_right hx.1, le_of_add_le_add_right hx.2⟩, hy⟩, -end +namespace finset +variables {α : Type*} +section preorder +variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] {a b : α} -lemma image_add_const_Ico (a b c : α) : (Ico a b).image ((+) c) = Ico (a + c) (b + c) := -begin - ext x, - rw [mem_image, mem_Ico], - split, - { rintro ⟨y, hy, hx⟩, - rw mem_Ico at hy, - rw [←hx, add_comm c], - exact ⟨add_le_add_right hy.1 c, add_lt_add_right hy.2 c⟩ }, - intro hx, - obtain ⟨y, hy⟩ := exists_add_of_le hx.1, - rw [hy, add_right_comm] at hx, - rw [eq_comm, add_right_comm, add_comm] at hy, - exact ⟨a + y, mem_Ico.2 ⟨le_of_add_le_add_right hx.1, lt_of_add_lt_add_right hx.2⟩, hy⟩, -end +@[simp] lemma nonempty_Icc : (Icc a b).nonempty ↔ a ≤ b := +by rw [←coe_nonempty, coe_Icc, set.nonempty_Icc] -lemma image_add_const_Ioc (a b c : α) : (Ioc a b).image ((+) c) = Ioc (a + c) (b + c) := -begin - ext x, - rw [mem_image, mem_Ioc], - split, - { rintro ⟨y, hy, hx⟩, - rw mem_Ioc at hy, - rw [←hx, add_comm c], - exact ⟨add_lt_add_right hy.1 c, add_le_add_right hy.2 c⟩ }, - intro hx, - obtain ⟨y, hy⟩ := exists_add_of_le hx.1.le, - rw [hy, add_right_comm] at hx, - rw [eq_comm, add_right_comm, add_comm] at hy, - exact ⟨a + y, mem_Ioc.2 ⟨lt_of_add_lt_add_right hx.1, le_of_add_le_add_right hx.2⟩, hy⟩, -end +@[simp] lemma nonempty_Ico : (Ico a b).nonempty ↔ a < b := +by rw [←coe_nonempty, coe_Ico, set.nonempty_Ico] -lemma image_add_const_Ioo (a b c : α) : (Ioo a b).image ((+) c) = Ioo (a + c) (b + c) := +@[simp] lemma nonempty_Ioc : (Ioc a b).nonempty ↔ a < b := +by rw [←coe_nonempty, coe_Ioc, set.nonempty_Ioc] + +@[simp] lemma nonempty_Ioo [densely_ordered α] : (Ioo a b).nonempty ↔ a < b := +by rw [←coe_nonempty, coe_Ioo, set.nonempty_Ioo] + +@[simp] lemma Icc_eq_empty_iff : Icc a b = ∅ ↔ ¬a ≤ b := +by rw [←coe_eq_empty, coe_Icc, set.Icc_eq_empty_iff] + +@[simp] lemma Ico_eq_empty_iff : Ico a b = ∅ ↔ ¬a < b := +by rw [←coe_eq_empty, coe_Ico, set.Ico_eq_empty_iff] + +@[simp] lemma Ioc_eq_empty_iff : Ioc a b = ∅ ↔ ¬a < b := +by rw [←coe_eq_empty, coe_Ioc, set.Ioc_eq_empty_iff] + +@[simp] lemma Ioo_eq_empty_iff [densely_ordered α] : Ioo a b = ∅ ↔ ¬a < b := +by rw [←coe_eq_empty, coe_Ioo, set.Ioo_eq_empty_iff] + +alias Icc_eq_empty_iff ↔ _ Icc_eq_empty +alias Ico_eq_empty_iff ↔ _ Ico_eq_empty +alias Ioc_eq_empty_iff ↔ _ Ioc_eq_empty + +@[simp] lemma Ioo_eq_empty (h : ¬a < b) : Ioo a b = ∅ := +eq_empty_iff_forall_not_mem.2 $ λ x hx, h ((mem_Ioo.1 hx).1.trans (mem_Ioo.1 hx).2) + +@[simp] lemma Icc_eq_empty_of_lt (h : b < a) : Icc a b = ∅ := +Icc_eq_empty h.not_le + +@[simp] lemma Ico_eq_empty_of_le (h : b ≤ a) : Ico a b = ∅ := +Ico_eq_empty h.not_lt + +@[simp] lemma Ioc_eq_empty_of_le (h : b ≤ a) : Ioc a b = ∅ := +Ioc_eq_empty h.not_lt + +@[simp] lemma Ioo_eq_empty_of_le (h : b ≤ a) : Ioo a b = ∅ := +Ioo_eq_empty h.not_lt + +variables (a b) + +@[simp] lemma Ico_self : Ico a a = ∅ := +by rw [←coe_eq_empty, coe_Ico, set.Ico_self] + +@[simp] lemma Ioc_self : Ioc a a = ∅ := +by rw [←coe_eq_empty, coe_Ioc, set.Ioc_self] + +@[simp] lemma Ioo_self : Ioo a a = ∅ := +by rw [←coe_eq_empty, coe_Ioo, set.Ioo_self] + +@[simp] lemma right_not_mem_Ico {a b : α} : b ∉ Ico a b := +by { rw [mem_Ico, not_and], exact λ _, lt_irrefl _ } + +lemma Ico_filter_lt_of_le_left [decidable_rel ((<) : α → α → Prop)] {a b c : α} (hca : c ≤ a) : + (Ico a b).filter (λ x, x < c) = ∅ := +finset.filter_false_of_mem (λ x hx, (hca.trans (mem_Ico.1 hx).1).not_lt) + +lemma Ico_filter_lt_of_right_le [decidable_rel ((<) : α → α → Prop)] {a b c : α} (hbc : b ≤ c) : + (Ico a b).filter (λ x, x < c) = Ico a b := +finset.filter_true_of_mem (λ x hx, (mem_Ico.1 hx).2.trans_le hbc) + +lemma Ico_filter_lt_of_le_right [decidable_rel ((<) : α → α → Prop)] {a b c : α} (hcb : c ≤ b) : + (Ico a b).filter (λ x, x < c) = Ico a c := begin ext x, - rw [mem_image, mem_Ioo], - split, - { rintro ⟨y, hy, hx⟩, - rw mem_Ioo at hy, - rw [←hx, add_comm c], - exact ⟨add_lt_add_right hy.1 c, add_lt_add_right hy.2 c⟩ }, - intro hx, - obtain ⟨y, hy⟩ := exists_add_of_le hx.1.le, - rw [hy, add_right_comm] at hx, - rw [eq_comm, add_right_comm, add_comm] at hy, - exact ⟨a + y, mem_Ioo.2 ⟨lt_of_add_lt_add_right hx.1, lt_of_add_lt_add_right hx.2⟩, hy⟩, + rw [mem_filter, mem_Ico, mem_Ico, and.right_comm], + exact and_iff_left_of_imp (λ h, h.2.trans_le hcb), end -lemma image_sub (a b c : α) (h : a ≤ c) : (Ico a b).image (λ x, x - c) = Ico (a - c) (b - c) := -begin - dsimp [image], - rw [multiset.Ico.map_sub _ _ _ h, ←multiset.to_finset_eq], - refl, -end +lemma Ico_filter_le_of_le_left {a b c : α} (hca : c ≤ a) : + (Ico a b).filter (λ x, c ≤ x) = Ico a b := +finset.filter_true_of_mem (λ x hx, hca.trans (mem_Ico.1 hx).1) + +lemma Ico_filter_le_of_right_le {a b c : α} : (Ico a b).filter (λ x, b ≤ x) = ∅ := +finset.filter_false_of_mem (λ x hx, (mem_Ico.1 hx).2.not_le) -lemma Ico_eq_range' (a b : ℕ) : Ico a b = (list.range' a (b - a)).to_finset := +lemma Ico_filter_le_of_left_le {a b c : α} (hac : a ≤ c) : + (Ico a b).filter (λ x, c ≤ x) = Ico c b := begin ext x, - rw [mem_Ico, list.mem_to_finset, list.mem_range'], - cases le_total a b, - { rw nat.add_sub_cancel' h }, - rw [nat.sub_eq_zero_iff_le.2 h, add_zero], - exact iff_of_false (λ hx, (h.trans hx.1).not_lt hx.2) (λ hx, hx.1.not_lt hx.2), + rw [mem_filter, mem_Ico, mem_Ico, and_comm, and.left_comm], + exact and_iff_right_of_imp (λ h, hac.trans h.1), end --- TODO: Once we have `has_lt_iff_add_one_le`, we can generalise that -/-- Currently only for ℕ -/ -@[simp] lemma card_Ico (a b : ℕ) : (Ico a b).card = a - b := -multiset.Ico.card _ _ +end preorder -theorem eq_empty_of_le {n m : ℕ} (h : m ≤ n) : Ico a b = ∅ := -eq_of_veq $ multiset.Ico.eq_zero_of_le h +section partial_order +variables [partial_order α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] {a b : α} -@[simp] theorem self_eq_empty (n : ℕ) : Ico n n = ∅ := -eq_empty_of_le $ le_refl n -@[simp] theorem Ico_nonempty_iff {a b : α} : (Ico a b).nonempty ↔ a < b := -begin - rw ←coe_nonempty, - rw set.Ico_eq, - split, - { - rintro ⟨x, hx⟩, - rw mem_Ico at hx, - } -end +@[simp] lemma Icc_self (a : α) : Icc a a = {a} := +by rw [←coe_eq_singleton, coe_Icc, set.Icc_self] -class has_lt_iff_add_one_le (α : Type*) [preorder α] [has_add α] [has_one α] : - Prop := -(lt_iff_add_one_le : ∀ a b : α, a < b ↔ a + 1 ≤ b) +lemma Ico_insert_right [decidable_eq α] (h : a ≤ b) : insert b (Ico a b) = Icc a b := +by rw [←coe_inj, coe_insert, coe_Icc, coe_Ico, set.insert_eq, set.union_comm, set.Ico_union_right h] -@[simp] theorem Ico_eq_empty_iff {a b : α} : Ico a b = ∅ ↔ ¬a < b := -by { rw [←not_nonempty_iff_eq_empty, not_iff_not], exact Ico_nonempty_iff } +lemma Ioo_insert_left [decidable_eq α] (h : a < b) : insert a (Ioo a b) = Ico a b := +by rw [←coe_inj, coe_insert, coe_Ioo, coe_Ico, set.insert_eq, set.union_comm, set.Ioo_union_left h] -theorem Ico_subset_Ico_iff {a₁ b₁ a₂ b₂ : α} (h : a₁ < b₁) : - Ico a₁ b₁ ⊆ Ico a₂ b₂ ↔ (a₂ ≤ a₁ ∧ b₁ ≤ b₂) := +@[simp] lemma Ico_inter_Ico_consecutive [decidable_eq α] (a b c : α) : Ico a b ∩ Ico b c = ∅ := begin - simp_rw [subset_iff, mem_Ico], - refine ⟨λ hx, ⟨(hx ⟨le_rfl, h⟩).1, _⟩, _⟩, - { exact (hx ⟨le_rfl, h⟩).1 }, - { refine le_of_pred_lt (@h (pred n₁) ⟨le_pred_of_lt hmn, pred_lt _⟩).2, - exact ne_of_gt (lt_of_le_of_lt (nat.zero_le m₁) hmn) }, - { rintros ⟨hm, hn⟩ k ⟨hmk, hkn⟩, - exact ⟨le_trans hm hmk, lt_of_lt_of_le hkn hn⟩ } + refine eq_empty_of_forall_not_mem (λ x hx, _), + rw [mem_inter, mem_Ico, mem_Ico] at hx, + exact hx.1.2.not_le hx.2.1, end -end ordered_cancel_add_comm_monoid +lemma Ico_disjoint_Ico_consecutive [decidable_eq α] (a b c : α) : disjoint (Ico a b) (Ico b c) := +le_of_eq $ Ico_inter_Ico_consecutive a b c -section ordered_semiring -variables {α : Type*} [decidable_eq α] [ordered_semiring α] [locally_finite_order α] +lemma Ico_filter_le_left {a b : α} (hab : a < b) : (Ico a b).filter (λ x, x ≤ a) = {a} := +begin + ext x, + rw [mem_filter, mem_Ico, mem_singleton, and.right_comm, ←le_antisymm_iff, eq_comm], + exact and_iff_left_of_imp (λ h, h.le.trans_lt hab), +end -end ordered_semiring +end partial_order section linear_order -variables {α : Type*} [decidable_eq α] [linear_order α] [locally_finite_order α] +variables [linear_order α] [locally_finite_order α] {a b : α} + +lemma Ico_subset_Ico_iff {a₁ b₁ a₂ b₂ : α} (h : a₁ < b₁) : + Ico a₁ b₁ ⊆ Ico a₂ b₂ ↔ a₂ ≤ a₁ ∧ b₁ ≤ b₂ := +by rw [←coe_subset, coe_Ico, coe_Ico, set.Ico_subset_Ico_iff h] lemma Ico_union_Ico_eq_Ico {a b c : α} (hab : a ≤ b) (hbc : b ≤ c) : Ico a b ∪ Ico b c = Ico a c := -begin - ext x, - rw [mem_union, ←mem_coe, ←mem_coe, ←mem_coe, coe_Ico, coe_Ico, coe_Ico, ←set.mem_union, - set.Ico_union_Ico_eq_Ico hab hbc], -end +by rw [←coe_inj, coe_union, coe_Ico, coe_Ico, coe_Ico, set.Ico_union_Ico_eq_Ico hab hbc] lemma Ico_union_Ico' {a b c d : α} (hcb : c ≤ b) (had : a ≤ d) : - Ico a b ∪ Ico c d = Ico (min n l) (max m k) := -by simp [←coe_inj, set.Ico_union_Ico' hlm hnk] + Ico a b ∪ Ico c d = Ico (min a c) (max b d) := +by rw [←coe_inj, coe_union, coe_Ico, coe_Ico, coe_Ico, set.Ico_union_Ico' hcb had] -lemma union {a b c d : α} (h₁ : min a b ≤ max c d) (h₂ : min c d ≤ max a b) : - Ico a b ∪ Ico l k = Ico (min n l) (max m k) := -by simp [←coe_inj, set.Ico_union_Ico h₁ h₂] +lemma Ico_union_Ico {a b c d : α} (h₁ : min a b ≤ max c d) (h₂ : min c d ≤ max a b) : + Ico a b ∪ Ico c d = Ico (min a c) (max b d) := +by rw [←coe_inj, coe_union, coe_Ico, coe_Ico, coe_Ico, set.Ico_union_Ico h₁ h₂] -@[simp] lemma inter_consecutive (n m l : ℕ) : Ico a b ∩ Ico m l = ∅ := +lemma Ico_inter_Ico {a b c d : α} : Ico a b ∩ Ico c d = Ico (max a c) (min b d) := +by rw [←coe_inj, coe_inter, coe_Ico, coe_Ico, coe_Ico, ←inf_eq_min, ←sup_eq_max, set.Ico_inter_Ico] + +@[simp] lemma Ico_filter_lt (a b c : α) : (Ico a b).filter (λ x, x < c) = Ico a (min b c) := begin - rw [← to_finset, ← to_finset, ← multiset.to_finset_inter, multiset.Ico.inter_consecutive], - simp, + cases le_total b c, + { rw [Ico_filter_lt_of_right_le h, min_eq_left h] }, + { rw [Ico_filter_lt_of_le_right h, min_eq_right h] } end -lemma inter {n m l k : ℕ} : Ico a b ∩ Ico l k = Ico (max n l) (min m k) := -by simp [←coe_inj, ←inf_eq_min, ←sup_eq_max, set.Ico_inter_Ico] - -lemma disjoint_consecutive (a b l : ℕ) : disjoint (Ico a b) (Ico m l) := -le_of_eq $ inter_consecutive n m l - -@[simp] theorem succ_singleton (n : ℕ) : Ico n (n+1) = {n} := -eq_of_veq $ multiset.Ico.succ_singleton - -theorem succ_top {n m : ℕ} (h : n ≤ m) : Ico n (m + 1) = insert m (Ico a b) := -by rw [← to_finset, multiset.Ico.succ_top h, multiset.to_finset_cons, to_finset] +@[simp] lemma Ico_filter_le (a b c : α) : (Ico a b).filter (λ x, c ≤ x) = Ico (max a c) b := +begin + cases le_total a c, + { rw [Ico_filter_le_of_left_le h, max_eq_right h] }, + { rw [Ico_filter_le_of_le_left h, max_eq_left h] } +end -theorem succ_top' {n m : ℕ} (h : n < m) : Ico a b = insert (m - 1) (Ico n (m - 1)) := +@[simp] lemma Ico_diff_Ico_left (a b c : α) : (Ico a b) \ (Ico a c) = Ico (max a c) b := begin - have w : m = m - 1 + 1 := (nat.sub_add_cancel (nat.one_le_of_lt h)).symm, - conv { to_lhs, rw w }, - rw succ_top, - exact nat.le_pred_of_lt h + cases le_total a c, + { ext x, + rw [mem_sdiff, mem_Ico, mem_Ico, mem_Ico, max_eq_right h, and.right_comm, not_and, not_lt], + exact and_congr_left' ⟨λ hx, hx.2 hx.1, λ hx, ⟨h.trans hx, λ _, hx⟩⟩ }, + { rw [Ico_eq_empty_of_le h, sdiff_empty, max_eq_left h] } end -theorem insert_succ_bot {n m : ℕ} (h : n < m) : insert n (Ico (n + 1) m) = Ico a b := -by rw [eq_comm, ← to_finset, multiset.Ico.eq_cons h, multiset.to_finset_cons, to_finset] +@[simp] lemma Ico_diff_Ico_right (a b c : α) : (Ico a b) \ (Ico c b) = Ico a (min b c) := +begin + cases le_total b c, + { rw [Ico_eq_empty_of_le h, sdiff_empty, min_eq_left h] }, + { ext x, + rw [mem_sdiff, mem_Ico, mem_Ico, mem_Ico, min_eq_right h, and_assoc, not_and', not_le], + exact and_congr_right' ⟨λ hx, hx.2 hx.1, λ hx, ⟨hx.trans_le h, λ _, hx⟩⟩ } +end -@[simp] theorem pred_singleton {m : ℕ} (h : 0 < m) : Ico (m - 1) m = {m - 1} := -eq_of_veq $ multiset.Ico.pred_singleton h +end linear_order -@[simp] theorem right_not_mem_Ico {a b : α} : b ∉ Ico a b := -by { rw [mem_Ico, not_and], exact λ _, lt_irrefl _ } +section ordered_cancel_add_comm_monoid +variables [ordered_cancel_add_comm_monoid α] [has_exists_add_of_le α] [decidable_eq α] + [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] -lemma filter_lt_of_top_le {n m l : ℕ} (hml : m ≤ l) : (Ico a b).filter (λ x, x < l) = Ico a b := -eq_of_veq $ multiset.Ico.filter_lt_of_top_le hml +lemma image_add_const_Icc (a b c : α) : (Icc a b).image ((+) c) = Icc (a + c) (b + c) := +begin + ext x, + rw [mem_image, mem_Icc], + split, + { rintro ⟨y, hy, rfl⟩, + rw mem_Icc at hy, + rw add_comm c, + exact ⟨add_le_add_right hy.1 c, add_le_add_right hy.2 c⟩ }, + { intro hx, + obtain ⟨y, hy⟩ := exists_add_of_le hx.1, + rw [hy, add_right_comm] at hx, + rw [eq_comm, add_right_comm, add_comm] at hy, + exact ⟨a + y, mem_Icc.2 ⟨le_of_add_le_add_right hx.1, le_of_add_le_add_right hx.2⟩, hy⟩ } +end -lemma filter_lt_of_le_bot {n m l : ℕ} (hln : l ≤ n) : (Ico a b).filter (λ x, x < l) = ∅ := -eq_of_veq $ multiset.Ico.filter_lt_of_le_bot hln +lemma image_add_const_Ico (a b c : α) : (Ico a b).image ((+) c) = Ico (a + c) (b + c) := +begin + ext x, + rw [mem_image, mem_Ico], + split, + { rintro ⟨y, hy, hx⟩, + rw mem_Ico at hy, + rw [←hx, add_comm c], + exact ⟨add_le_add_right hy.1 c, add_lt_add_right hy.2 c⟩ }, + { intro hx, + obtain ⟨y, hy⟩ := exists_add_of_le hx.1, + rw [hy, add_right_comm] at hx, + rw [eq_comm, add_right_comm, add_comm] at hy, + exact ⟨a + y, mem_Ico.2 ⟨le_of_add_le_add_right hx.1, lt_of_add_lt_add_right hx.2⟩, hy⟩ } +end -lemma filter_Ico_bot {n m : ℕ} (hnm : n < m) : (Ico a b).filter (λ x, x ≤ n) = {n} := -eq_of_veq $ multiset.Ico.filter_le_of_bot hnm +lemma image_add_const_Ioc (a b c : α) : (Ioc a b).image ((+) c) = Ioc (a + c) (b + c) := +begin + ext x, + rw [mem_image, mem_Ioc], + split, + { rintro ⟨y, hy, rfl⟩, + rw mem_Ioc at hy, + rw add_comm c, + exact ⟨add_lt_add_right hy.1 c, add_le_add_right hy.2 c⟩ }, + { intro hx, + obtain ⟨y, hy⟩ := exists_add_of_le hx.1.le, + rw [hy, add_right_comm] at hx, + rw [eq_comm, add_right_comm, add_comm] at hy, + exact ⟨a + y, mem_Ioc.2 ⟨lt_of_add_lt_add_right hx.1, le_of_add_le_add_right hx.2⟩, hy⟩ } +end -lemma filter_lt_of_ge {n m l : ℕ} (hlm : l ≤ m) : (Ico a b).filter (λ x, x < l) = Ico n l := -eq_of_veq $ multiset.Ico.filter_lt_of_ge hlm +lemma image_add_const_Ioo (a b c : α) : (Ioo a b).image ((+) c) = Ioo (a + c) (b + c) := +begin + ext x, + rw [mem_image, mem_Ioo], + split, + { rintro ⟨y, hy, rfl⟩, + rw mem_Ioo at hy, + rw add_comm c, + exact ⟨add_lt_add_right hy.1 c, add_lt_add_right hy.2 c⟩ }, + { intro hx, + obtain ⟨y, hy⟩ := exists_add_of_le hx.1.le, + rw [hy, add_right_comm] at hx, + rw [eq_comm, add_right_comm, add_comm] at hy, + exact ⟨a + y, mem_Ioo.2 ⟨lt_of_add_lt_add_right hx.1, lt_of_add_lt_add_right hx.2⟩, hy⟩ } +end -@[simp] lemma filter_lt (n m l : ℕ) : (Ico a b).filter (λ x, x < l) = Ico n (min m l) := -eq_of_veq $ multiset.Ico.filter_lt n m l +end ordered_cancel_add_comm_monoid -lemma filter_le_of_le_bot {n m l : ℕ} (hln : l ≤ n) : (Ico a b).filter (λ x, l ≤ x) = Ico a b := -eq_of_veq $ multiset.Ico.filter_le_of_le_bot hln +section canonically_ordered_add_monoid +variables [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] + [contravariant_class α α (+) (≤)] [decidable_eq α] [decidable_rel ((≤) : α → α → Prop)] + [locally_finite_order α] -lemma filter_le_of_top_le {n m l : ℕ} (hml : m ≤ l) : (Ico a b).filter (λ x, l ≤ x) = ∅ := -eq_of_veq $ multiset.Ico.filter_le_of_top_le hml +lemma image_sub_const_Ico (a b c : α) (h : c ≤ a) : + (Ico a b).image (λ x, x - c) = Ico (a - c) (b - c) := +begin + ext x, + rw [mem_image], + split, + { rintro ⟨x, hx, rfl⟩, + rw [mem_Ico] at ⊢ hx, + exact ⟨sub_le_sub_right' hx.1 _, sub_lt_sub_right_of_le (h.trans hx.1) hx.2⟩ }, + rw [mem_image, mem_Ico], + rw [], + dsimp [image], + rw [multiset.Ico.map_sub _ _ _ h, ←multiset.to_finset_eq], + refl, +end -lemma filter_le_of_le {n m l : ℕ} (hnl : n ≤ l) : (Ico a b).filter (λ x, l ≤ x) = Ico l m := -eq_of_veq $ multiset.Ico.filter_le_of_le hnl +end ordered_cancel_add_comm_monoid -@[simp] lemma filter_le (n m l : ℕ) : (Ico a b).filter (λ x, l ≤ x) = Ico (max n l) m := -eq_of_veq $ multiset.Ico.filter_le n m l +section ordered_semiring +variables [decidable_eq α] [ordered_semiring α] [locally_finite_order α] -@[simp] lemma diff_left (l n m : ℕ) : (Ico a b) \ (Ico n l) = Ico (max n l) m := -by ext k; by_cases n ≤ k; simp [h, and_comm] +end ordered_semiring -@[simp] lemma diff_right (l n m : ℕ) : (Ico a b) \ (Ico l m) = Ico n (min m l) := -have ∀k, (k < m ∧ (l ≤ k → m ≤ k)) ↔ (k < m ∧ k < l) := - assume k, and_congr_right $ assume hk, by rw [← not_imp_not]; simp [hk], -by ext k; by_cases n ≤ k; simp [h, this] +section linear_order +variables [linear_order α] [locally_finite_order α] -lemma image_const_sub {k m n : ℕ} (hkn : k ≤ n) : - (Ico k m).image (λ j, n - j) = Ico (n + 1 - m) (n + 1 - k) := +lemma Ico_image_const_sub {a b c : α} : + (Ico a b).image (λ x, c - x) = Ioc (c - a) (c - b) := begin rw [nat.sub_add_comm hkn], ext j, @@ -258,15 +331,5 @@ begin { exact nat.sub_sub_self hj } } end -lemma range_eq_Ico (n : ℕ) : finset.range n = finset.Ico 0 n := -by { ext i, simp } - -lemma range_image_pred_top_sub (n : ℕ) : - (finset.range n).image (λ j, n - 1 - j) = finset.range n := -begin - cases n, - { simp }, - { simp [range_eq_Ico, Ico.image_const_sub] } -end - +end linear_order end finset diff --git a/src/data/finset/mwe.lean b/src/data/finset/mwe.lean new file mode 100644 index 0000000000000..e69de29bb2d1d diff --git a/src/data/nat/intervals.lean b/src/data/nat/intervals.lean index 878bc016596cb..9b21f66f97ddc 100644 --- a/src/data/nat/intervals.lean +++ b/src/data/nat/intervals.lean @@ -3,8 +3,7 @@ Copyright (c) 2021 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import data.nat.basic -import order.locally_finite +import data.finset.intervals /-! # Intervals of naturals @@ -52,29 +51,39 @@ instance : locally_finite_order ℕ := namespace nat variables (a b : ℕ) -@[simp] lemma card_finset_Icc : (Icc a b).card = b + 1 - a := +lemma Icc_eq_range' : Icc a b = (list.range' a (b + 1 - a)).to_finset := rfl +lemma Ico_eq_range' : Ico a b = (list.range' a (b - a)).to_finset := rfl +lemma Ioc_eq_range' : Ioc a b = (list.range' (a + 1) (b - a)).to_finset := rfl +lemma Ioo_eq_range' : Ioo a b = (list.range' (a + 1) (b - a - 1)).to_finset := rfl + +lemma Iio_eq_range : Iio = range := +by { ext b x, rw [mem_Iio, mem_range] } + +lemma Ico_zero_eq_range : Ico 0 = range := +by rw [←bot_eq_zero, ←Iio_eq_Ico, Iio_eq_range] + +lemma _root_.finset.range_eq_Ico : range = Ico 0 := Ico_zero_eq_range.symm + +lemma range_image_pred_top_sub (n : ℕ) : + (finset.range n).image (λ j, n - 1 - j) = finset.range n := begin - change (list.to_finset _).card = _, - rw [list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'], + rw range_eq_Ico, + -- cases n, + -- { simp }, + -- { simp [range_eq_Ico, Ico.image_const_sub] } end +@[simp] lemma card_finset_Icc : (Icc a b).card = b + 1 - a := +by rw [Icc_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] + @[simp] lemma card_finset_Ico : (Ico a b).card = b - a := -begin - change (list.to_finset _).card = _, - rw [list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'], -end +by rw [Ico_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] @[simp] lemma card_finset_Ioc : (Ioc a b).card = b - a := -begin - change (list.to_finset _).card = _, - rw [list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'], -end +by rw [Ioc_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] @[simp] lemma card_finset_Ioo : (Ioo a b).card = b - a - 1 := -begin - change (list.to_finset _).card = _, - rw [list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'], -end +by rw [Ioo_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := by rw [←card_finset_Icc, fintype.card_of_finset] @@ -88,4 +97,39 @@ by rw [←card_finset_Ioc, fintype.card_of_finset] @[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := by rw [←card_finset_Ioo, fintype.card_of_finset] +-- TODO@Yaël: Generalize all the following lemmas to `succ_order` + +lemma Icc_succ_left : Icc a.succ b = Ioc a b := +by { ext x, rw [mem_Icc, mem_Ioc, succ_le_iff] } + +lemma Ico_succ_right : Ico a b.succ = Icc a b := +by { ext x, rw [mem_Ico, mem_Icc, lt_succ_iff] } + +lemma Ico_succ_left : Ico a.succ b = Ioo a b := +by { ext x, rw [mem_Ico, mem_Ioo, succ_le_iff] } + +lemma Icc_pred_right {b : ℕ} (h : 0 < b) : Icc a (b - 1) = Ico a b := +by { ext x, rw [mem_Icc, mem_Ico, lt_iff_le_pred h] } + +lemma Ioc_succ_succ : Ioc a.succ b.succ = Ico a b := +by { ext x, rw [mem_Ioc, mem_Ico, succ_le_iff, lt_succ_iff] } + +@[simp] lemma Ico_succ_singleton : Ico a (a + 1) = {a} := +by rw [Ico_succ_right, Icc_self] + +@[simp] lemma Ico_pred_singleton {a : ℕ} (h : 0 < a) : Ico (a - 1) a = {a - 1} := +by rw [←Icc_pred_right _ h, Icc_self] + +variables {a b} + +lemma Ico_succ_right_eq_insert_Ico (h : a ≤ b) : Ico a (b + 1) = insert b (Ico a b) := +by rw [Ico_succ_right, ←Ico_insert_right h] + +lemma Ico_insert_succ_left (h : a < b) : insert a (Ico a.succ b) = Ico a b := +by rw [Ico_succ_left, ←Ioo_insert_left h] + +lemma Ico_image_const_sub_eq_Ico {k m n : α} (hkn : k ≤ n) : + (Ico k m).image (λ j, n - j) = Ico (n + 1 - m) (n + 1 - k) := +by rw [Ico_image_const_sub, Ioc_succ_succ] + end nat diff --git a/src/data/nat/totient.lean b/src/data/nat/totient.lean index 3fdfe8b114798..9a69fbb6ee1b8 100644 --- a/src/data/nat/totient.lean +++ b/src/data/nat/totient.lean @@ -183,7 +183,7 @@ begin exact one_ne_zero h } }, have hsplit : finset.range p = {0} ∪ (finset.Ico 1 p), { rw [finset.range_eq_Ico, ←finset.Ico.union_consecutive zero_le_one hp.le, - finset.Ico.succ_singleton] }, + finset.Ico_succ_singleton] }, have hempty : finset.filter p.coprime {0} = ∅, { simp only [finset.filter_singleton, nat.coprime_zero_right, hp.ne', if_false] }, rw [totient_eq_card_coprime, hsplit, finset.filter_union, hempty, finset.empty_union, diff --git a/src/measure_theory/decomposition/unsigned_hahn.lean b/src/measure_theory/decomposition/unsigned_hahn.lean index 16680fd974507..1845327959e89 100644 --- a/src/measure_theory/decomposition/unsigned_hahn.lean +++ b/src/measure_theory/decomposition/unsigned_hahn.lean @@ -123,7 +123,7 @@ begin { assume n m hnm, have : n ≤ m + 1 := le_of_lt (nat.succ_le_succ hnm), simp only [f], - rw [finset.Ico.succ_top this, finset.inf_insert, set.inter_comm], + rw [finset.Ico_insert_right this, finset.inf_insert, set.inter_comm], refl }, have le_d_f : ∀n m, m ≤ n → γ - 2 * ((1 / 2) ^ m) + (1 / 2) ^ n ≤ d (f m n), @@ -131,7 +131,7 @@ begin refine nat.le_induction _ _ n h, { have := he₂ m, simp only [f], - rw [finset.Ico.succ_singleton, finset.inf_singleton], + rw [finset.Ico_succ_singleton, finset.inf_singleton], exact aux this }, { assume n (hmn : m ≤ n) ih, have : γ + (γ - 2 * (1 / 2)^m + (1 / 2) ^ (n + 1)) ≤ γ + d (f m (n + 1)), diff --git a/src/number_theory/divisors.lean b/src/number_theory/divisors.lean index 7c95daa8c852c..35e0556c6fa81 100644 --- a/src/number_theory/divisors.lean +++ b/src/number_theory/divisors.lean @@ -70,7 +70,7 @@ end lemma divisors_eq_proper_divisors_insert_self_of_pos (h : 0 < n): divisors n = has_insert.insert n (proper_divisors n) := -by rw [divisors, proper_divisors, finset.Ico.succ_top h, finset.filter_insert, if_pos (dvd_refl n)] +by rw [divisors, proper_divisors, finset.Ico_insert_right h, finset.filter_insert, if_pos (dvd_refl n)] @[simp] lemma mem_divisors {m : ℕ} : diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 54db2f582df34..cdcd262dcf4f5 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -240,6 +240,9 @@ Icc a ⊤ def Ioi (a : α) : finset α := Ioc a ⊤ +lemma Ici_eq_Icc (a : α) : Ici a = Icc a ⊤ := rfl +lemma Ioi_eq_Ioc (a : α) : Ioi a = Ioc a ⊤ := rfl + @[simp, norm_cast] lemma coe_Ici (a : α) : (Ici a : set α) = set.Ici a := by rw [Ici, coe_Icc, set.Icc_top] @@ -265,6 +268,9 @@ Icc ⊥ b def Iio (b : α) : finset α := Ico ⊥ b +lemma Iic_eq_Icc : Iic = Icc (⊥ : α) := rfl +lemma Iio_eq_Ico : Iio = Ico (⊥ : α) := rfl + @[simp, norm_cast] lemma coe_Iic (b : α) : (Iic b : set α) = set.Iic b := by rw [Iic, coe_Icc, set.Icc_bot] diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index b9bf737740199..bb8e0c1e89302 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -196,7 +196,7 @@ begin calc dist (f m) (f (n+1)) ≤ dist (f m) (f n) + dist _ _ : dist_triangle _ _ _ ... ≤ ∑ i in finset.Ico m n, _ + _ : add_le_add hrec (le_refl _) ... = ∑ i in finset.Ico m (n+1), _ : - by rw [finset.Ico.succ_top hn, finset.sum_insert, add_comm]; simp } + by rw [finset.Ico_insert_right hn, finset.sum_insert, add_comm]; simp } end /-- The triangle (polygon) inequality for sequences of points; `finset.range` version. -/ diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index 3d36f5699c727..e0cd0e3c00c4d 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -137,7 +137,7 @@ begin calc edist (f m) (f (n+1)) ≤ edist (f m) (f n) + edist (f n) (f (n+1)) : edist_triangle _ _ _ ... ≤ ∑ i in finset.Ico m n, _ + _ : add_le_add hrec (le_refl _) ... = ∑ i in finset.Ico m (n+1), _ : - by rw [finset.Ico.succ_top hn, finset.sum_insert, add_comm]; simp } + by rw [finset.Ico_insert_right hn, finset.sum_insert, add_comm]; simp } end /-- The triangle (polygon) inequality for sequences of points; `finset.range` version. -/ From 19bb5ec1f37215419e467b92c2afd28739c7622a Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Thu, 16 Sep 2021 19:16:53 +0200 Subject: [PATCH 11/41] add some subtyping glue --- src/algebra/big_operators/intervals.lean | 12 +-- src/data/finset/intervals.lean | 53 ------------ src/data/int/intervals.lean | 38 +++++---- src/data/nat/intervals.lean | 87 ++++++++++++++------ src/data/nat/totient.lean | 4 +- src/data/pnat/intervals.lean | 63 +++++++------- src/data/polynomial/iterated_deriv.lean | 2 +- src/order/locally_finite.lean | 82 ++++++++++++++---- src/topology/metric_space/basic.lean | 6 +- src/topology/metric_space/emetric_space.lean | 6 +- 10 files changed, 197 insertions(+), 156 deletions(-) diff --git a/src/algebra/big_operators/intervals.lean b/src/algebra/big_operators/intervals.lean index f272fbc3a6bad..b436e5317a249 100644 --- a/src/algebra/big_operators/intervals.lean +++ b/src/algebra/big_operators/intervals.lean @@ -34,7 +34,7 @@ lemma prod_Ico_add (f : ℕ → β) (m n k : ℕ) : lemma sum_Ico_succ_top {δ : Type*} [add_comm_monoid δ] {a b : ℕ} (hab : a ≤ b) (f : ℕ → δ) : (∑ k in Ico a (b + 1), f k) = (∑ k in Ico a b, f k) + f b := -by rw [Ico_insert_right hab, sum_insert right_not_mem_Ico, add_comm] +by rw [nat.Ico_succ_right_eq_insert_Ico hab, sum_insert right_not_mem_Ico, add_comm] @[to_additive] lemma prod_Ico_succ_top {a b : ℕ} (hab : a ≤ b) (f : ℕ → β) : @@ -44,7 +44,7 @@ lemma prod_Ico_succ_top {a b : ℕ} (hab : a ≤ b) (f : ℕ → β) : lemma sum_eq_sum_Ico_succ_bot {δ : Type*} [add_comm_monoid δ] {a b : ℕ} (hab : a < b) (f : ℕ → δ) : (∑ k in Ico a b, f k) = f a + (∑ k in Ico (a + 1) b, f k) := have ha : a ∉ Ico (a + 1) b, by simp, -by rw [← sum_insert ha, Ico.insert_succ_bot hab] +by rw [← sum_insert ha, nat.Ico_insert_succ_left hab] @[to_additive] lemma prod_eq_prod_Ico_succ_bot {a b : ℕ} (hab : a < b) (f : ℕ → β) : @@ -54,12 +54,12 @@ lemma prod_eq_prod_Ico_succ_bot {a b : ℕ} (hab : a < b) (f : ℕ → β) : @[to_additive] lemma prod_Ico_consecutive (f : ℕ → β) {m n k : ℕ} (hmn : m ≤ n) (hnk : n ≤ k) : (∏ i in Ico m n, f i) * (∏ i in Ico n k, f i) = (∏ i in Ico m k, f i) := -Ico_union_Ico_consecutive hmn hnk ▸ eq.symm $ prod_union $ Ico.disjoint_consecutive m n k +Ico_union_Ico_eq_Ico hmn hnk ▸ eq.symm $ prod_union $ Ico_disjoint_Ico_consecutive m n k @[to_additive] lemma prod_range_mul_prod_Ico (f : ℕ → β) {m n : ℕ} (h : m ≤ n) : (∏ k in range m, f k) * (∏ k in Ico m n, f k) = (∏ k in range n, f k) := -Ico.zero_bot m ▸ Ico.zero_bot n ▸ prod_Ico_consecutive f (nat.zero_le m) h +m.Ico_zero_eq_range ▸ n.Ico_zero_eq_range ▸ prod_Ico_consecutive f m.zero_le h @[to_additive] lemma prod_Ico_eq_mul_inv {δ : Type*} [comm_group δ] (f : ℕ → δ) {m n : ℕ} (h : m ≤ n) : @@ -90,7 +90,7 @@ lemma prod_Ico_eq_prod_range (f : ℕ → β) (m n : ℕ) : (∏ k in Ico m n, f k) = (∏ k in range (n - m), f (m + k)) := begin by_cases h : m ≤ n, - { rw [← Ico.zero_bot, prod_Ico_add, zero_add, nat.sub_add_cancel h] }, + { rw [←nat.Ico_zero_eq_range, prod_Ico_add, zero_add, nat.sub_add_cancel h] }, { replace h : n ≤ m := le_of_not_ge h, rw [Ico_eq_empty_of_le h, nat.sub_eq_zero_of_le h, range_zero, prod_empty, prod_empty] } end @@ -102,7 +102,7 @@ begin { intros i hi, exact (add_le_add_iff_right 1).1 (le_trans (nat.lt_iff_add_one_le.1 hi) h) }, cases lt_or_le k m with hkm hkm, - { rw [← finset.Ico.image_const_sub (this _ hkm)], + { rw [← nat.Ico_image_const_sub_eq_Ico (this _ hkm)], refine (prod_image _).symm, simp only [mem_Ico], rintros i ⟨ki, im⟩ j ⟨kj, jm⟩ Hij, diff --git a/src/data/finset/intervals.lean b/src/data/finset/intervals.lean index ee9ae40ea3f78..60fbf36b4a340 100644 --- a/src/data/finset/intervals.lean +++ b/src/data/finset/intervals.lean @@ -279,57 +279,4 @@ begin end end ordered_cancel_add_comm_monoid - -section canonically_ordered_add_monoid -variables [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] - [contravariant_class α α (+) (≤)] [decidable_eq α] [decidable_rel ((≤) : α → α → Prop)] - [locally_finite_order α] - -lemma image_sub_const_Ico (a b c : α) (h : c ≤ a) : - (Ico a b).image (λ x, x - c) = Ico (a - c) (b - c) := -begin - ext x, - rw [mem_image], - split, - { rintro ⟨x, hx, rfl⟩, - rw [mem_Ico] at ⊢ hx, - exact ⟨sub_le_sub_right' hx.1 _, sub_lt_sub_right_of_le (h.trans hx.1) hx.2⟩ }, - rw [mem_image, mem_Ico], - rw [], - dsimp [image], - rw [multiset.Ico.map_sub _ _ _ h, ←multiset.to_finset_eq], - refl, -end - -end ordered_cancel_add_comm_monoid - -section ordered_semiring -variables [decidable_eq α] [ordered_semiring α] [locally_finite_order α] - -end ordered_semiring - -section linear_order -variables [linear_order α] [locally_finite_order α] - -lemma Ico_image_const_sub {a b c : α} : - (Ico a b).image (λ x, c - x) = Ioc (c - a) (c - b) := -begin - rw [nat.sub_add_comm hkn], - ext j, - simp only [mem, mem_image, exists_prop, nat.lt_iff_add_one_le, add_le_add_iff_right], - split, - { rintros ⟨j, ⟨hjk, hjm⟩, rfl⟩, - split, - { simp only [← nat.add_sub_add_right n 1 j, nat.sub_le_sub_left, hjm] }, - { exact nat.sub_le_sub_left _ hjk } }, - { rintros ⟨hm, hk⟩, - have hj : j ≤ n := le_trans hk (nat.sub_le_self _ _), - refine ⟨n - j, ⟨_, _⟩, _⟩, - { apply nat.le_sub_right_of_add_le, - rwa nat.le_sub_left_iff_add_le hkn at hk }, - { rwa [← nat.sub_add_comm hj, nat.sub_le_iff] }, - { exact nat.sub_sub_self hj } } -end - -end linear_order end finset diff --git a/src/data/int/intervals.lean b/src/data/int/intervals.lean index a840933c122a7..52e0f14c1b62e 100644 --- a/src/data/int/intervals.lean +++ b/src/data/int/intervals.lean @@ -14,18 +14,18 @@ intervals as finsets and fintypes. -/ open finset int +private lemma inj (a : ℤ) : function.injective (λ n : ℕ, (n : ℤ) + a) := +λ m n, by { rw [add_left_inj, coe_nat_inj'], exact id } instance : locally_finite_order ℤ := -have hinj : ∀ a : ℤ, function.injective (λ n : ℕ, (n : ℤ) + a), -from λ a m n, by { rw [add_left_inj, coe_nat_inj'], exact id }, { finset_Icc := λ a b, (finset.range (b + 1 - a).to_nat).map - ⟨λ n, n + a, hinj a⟩, + ⟨λ n, n + a, inj a⟩, finset_Ico := λ a b, (finset.range (b - a).to_nat).map - ⟨λ n, n + a, hinj a⟩, + ⟨λ n, n + a, inj a⟩, finset_Ioc := λ a b, (finset.range (b - a).to_nat).map - ⟨λ n, n + (a + 1), hinj (a + 1)⟩, + ⟨λ n, n + (a + 1), inj (a + 1)⟩, finset_Ioo := λ a b, (finset.range (b - a - 1).to_nat).map - ⟨λ n, n + (a + 1), hinj (a + 1)⟩, + ⟨λ n, n + (a + 1), inj (a + 1)⟩, finset_mem_Icc := λ a b x, begin simp only [int.lt_to_nat, exists_prop, mem_range, add_comm, function.embedding.coe_fn_mk, mem_map], @@ -76,30 +76,38 @@ from λ a m n, by { rw [add_left_inj, coe_nat_inj'], exact id }, end } namespace int -variables (a b : ℤ). +variables (a b : ℤ) -@[simp] lemma card_finset_Icc : (Icc a b).card = (b + 1 - a).to_nat := +lemma Icc_eq_finset_map : Icc a b = (finset.range (b + 1 - a).to_nat).map ⟨λ n, n + a, inj a⟩ := rfl +lemma Ico_eq_finset_map : Ico a b = (finset.range (b - a).to_nat).map ⟨λ n, n + a, inj a⟩ := rfl +lemma Ioc_eq_finset_map : Ioc a b = (finset.range (b - a).to_nat).map ⟨λ n, n + (a + 1), inj _⟩ := +rfl +lemma Ioo_eq_finset_map : Ioo a b = (finset.range (b - a - 1).to_nat).map + ⟨λ n, n + (a + 1), inj _⟩ := +rfl + +@[simp] lemma card_Icc : (Icc a b).card = (b + 1 - a).to_nat := by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } -@[simp] lemma card_finset_Ico : (Ico a b).card = (b - a).to_nat := +@[simp] lemma card_Ico : (Ico a b).card = (b - a).to_nat := by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } -@[simp] lemma card_finset_Ioc : (Ioc a b).card = (b - a).to_nat := +@[simp] lemma card_Ioc : (Ioc a b).card = (b - a).to_nat := by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } -@[simp] lemma card_finset_Ioo : (Ioo a b).card = (b - a - 1).to_nat := +@[simp] lemma card_Ioo : (Ioo a b).card = (b - a - 1).to_nat := by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = (b + 1 - a).to_nat := -by rw [←card_finset_Icc, fintype.card_of_finset] +by rw [←card_Icc, fintype.card_of_finset] @[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = (b - a).to_nat := -by rw [←card_finset_Ico, fintype.card_of_finset] +by rw [←card_Ico, fintype.card_of_finset] @[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = (b - a).to_nat := -by rw [←card_finset_Ioc, fintype.card_of_finset] +by rw [←card_Ioc, fintype.card_of_finset] @[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = (b - a - 1).to_nat := -by rw [←card_finset_Ioo, fintype.card_of_finset] +by rw [←card_Ioo, fintype.card_of_finset] end int diff --git a/src/data/nat/intervals.lean b/src/data/nat/intervals.lean index 9b21f66f97ddc..18c9e7c14f6d0 100644 --- a/src/data/nat/intervals.lean +++ b/src/data/nat/intervals.lean @@ -49,7 +49,7 @@ instance : locally_finite_order ℕ := end } namespace nat -variables (a b : ℕ) +variables (a b c : ℕ) lemma Icc_eq_range' : Icc a b = (list.range' a (b + 1 - a)).to_finset := rfl lemma Ico_eq_range' : Ico a b = (list.range' a (b - a)).to_finset := rfl @@ -59,43 +59,34 @@ lemma Ioo_eq_range' : Ioo a b = (list.range' (a + 1) (b - a - 1)).to_finset := r lemma Iio_eq_range : Iio = range := by { ext b x, rw [mem_Iio, mem_range] } -lemma Ico_zero_eq_range : Ico 0 = range := +lemma Ico_zero_eq_range : Ico 0 a = range a := by rw [←bot_eq_zero, ←Iio_eq_Ico, Iio_eq_range] -lemma _root_.finset.range_eq_Ico : range = Ico 0 := Ico_zero_eq_range.symm +lemma _root_.finset.range_eq_Ico : range = Ico 0 := by { funext, exact (Ico_zero_eq_range n).symm } -lemma range_image_pred_top_sub (n : ℕ) : - (finset.range n).image (λ j, n - 1 - j) = finset.range n := -begin - rw range_eq_Ico, - -- cases n, - -- { simp }, - -- { simp [range_eq_Ico, Ico.image_const_sub] } -end - -@[simp] lemma card_finset_Icc : (Icc a b).card = b + 1 - a := +@[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := by rw [Icc_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] -@[simp] lemma card_finset_Ico : (Ico a b).card = b - a := +@[simp] lemma card_Ico : (Ico a b).card = b - a := by rw [Ico_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] -@[simp] lemma card_finset_Ioc : (Ioc a b).card = b - a := +@[simp] lemma card_Ioc : (Ioc a b).card = b - a := by rw [Ioc_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] -@[simp] lemma card_finset_Ioo : (Ioo a b).card = b - a - 1 := +@[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := by rw [Ioo_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := -by rw [←card_finset_Icc, fintype.card_of_finset] +by rw [←card_Icc, fintype.card_of_finset] @[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = b - a := -by rw [←card_finset_Ico, fintype.card_of_finset] +by rw [←card_Ico, fintype.card_of_finset] @[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = b - a := -by rw [←card_finset_Ioc, fintype.card_of_finset] +by rw [←card_Ioc, fintype.card_of_finset] @[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := -by rw [←card_finset_Ioo, fintype.card_of_finset] +by rw [←card_Ioo, fintype.card_of_finset] -- TODO@Yaël: Generalize all the following lemmas to `succ_order` @@ -111,8 +102,8 @@ by { ext x, rw [mem_Ico, mem_Ioo, succ_le_iff] } lemma Icc_pred_right {b : ℕ} (h : 0 < b) : Icc a (b - 1) = Ico a b := by { ext x, rw [mem_Icc, mem_Ico, lt_iff_le_pred h] } -lemma Ioc_succ_succ : Ioc a.succ b.succ = Ico a b := -by { ext x, rw [mem_Ioc, mem_Ico, succ_le_iff, lt_succ_iff] } +lemma Ico_succ_succ : Ico a.succ b.succ = Ioc a b := +by { ext x, rw [mem_Ico, mem_Ioc, succ_le_iff, lt_succ_iff] } @[simp] lemma Ico_succ_singleton : Ico a (a + 1) = {a} := by rw [Ico_succ_right, Icc_self] @@ -120,7 +111,7 @@ by rw [Ico_succ_right, Icc_self] @[simp] lemma Ico_pred_singleton {a : ℕ} (h : 0 < a) : Ico (a - 1) a = {a - 1} := by rw [←Icc_pred_right _ h, Icc_self] -variables {a b} +variables {a b c} lemma Ico_succ_right_eq_insert_Ico (h : a ≤ b) : Ico a (b + 1) = insert b (Ico a b) := by rw [Ico_succ_right, ←Ico_insert_right h] @@ -128,8 +119,52 @@ by rw [Ico_succ_right, ←Ico_insert_right h] lemma Ico_insert_succ_left (h : a < b) : insert a (Ico a.succ b) = Ico a b := by rw [Ico_succ_left, ←Ioo_insert_left h] -lemma Ico_image_const_sub_eq_Ico {k m n : α} (hkn : k ≤ n) : - (Ico k m).image (λ j, n - j) = Ico (n + 1 - m) (n + 1 - k) := -by rw [Ico_image_const_sub, Ioc_succ_succ] +lemma image_sub_const_Ico (h : c ≤ a) : + (Ico a b).image (λ x, x - c) = Ico (a - c) (b - c) := +begin + ext x, + rw mem_image, + split, + { rintro ⟨x, hx, rfl⟩, + rw [mem_Ico] at ⊢ hx, + exact ⟨sub_le_sub_right' hx.1 _, sub_lt_sub_right_of_le (h.trans hx.1) hx.2⟩ }, + { rintro h, + refine ⟨x + c, _, add_sub_cancel_right⟩, + rw mem_Ico at ⊢ h, + exact ⟨sub_le_iff_right.1 h.1, nat.add_lt_of_lt_sub_right h.2⟩ } +end + +lemma Ico_image_const_sub_eq_Ico (hac : a ≤ c) : + (Ico a b).image (λ x, c - x) = Ico (c + 1 - b) (c + 1 - a) := +begin + ext x, + rw [mem_image, mem_Ico], + split, + { rintro ⟨x, hx, rfl⟩, + rw mem_Ico at hx, + refine ⟨ _, ((nat.sub_le_sub_left_iff hac).2 hx.1).trans_lt ((nat.sub_lt_sub_right_iff hac).2 (nat.lt_succ_self _))⟩, + cases lt_or_le c b, + { rw nat.sub_eq_zero_of_le h, + exact zero_le _ }, + { rw ←succ_sub_succ c, + exact (nat.sub_le_sub_left_iff (succ_le_succ $ hx.2.le.trans h)).2 hx.2 } }, + { rintro ⟨hb, ha⟩, + rw [nat.lt_sub_left_iff_add_lt, lt_succ_iff] at ha, + have hx : x ≤ c := (nat.le_add_left _ _).trans ha, + refine ⟨c - x, _, nat.sub_sub_self hx⟩, + { rw mem_Ico, + refine ⟨nat.le_sub_right_of_add_le ha, _⟩, + rw [nat.sub_lt_left_iff_lt_add hx, ←succ_le_iff], + exact nat.le_add_of_sub_le_right hb } } +end + +lemma range_image_pred_top_sub (n : ℕ) : + (finset.range n).image (λ j, n - 1 - j) = finset.range n := +begin + cases n, + { rw [range_zero, image_empty] }, + { rw [finset.range_eq_Ico, Ico_image_const_sub_eq_Ico (zero_le _)], + simp_rw [succ_sub_succ, nat.sub_zero, nat.sub_self] } +end end nat diff --git a/src/data/nat/totient.lean b/src/data/nat/totient.lean index 9a69fbb6ee1b8..3027ee6308b18 100644 --- a/src/data/nat/totient.lean +++ b/src/data/nat/totient.lean @@ -182,12 +182,12 @@ begin rw [totient_one, nat.sub_self] at h, exact one_ne_zero h } }, have hsplit : finset.range p = {0} ∪ (finset.Ico 1 p), - { rw [finset.range_eq_Ico, ←finset.Ico.union_consecutive zero_le_one hp.le, + { rw [finset.range_eq_Ico, ←finset.Ico_union_Ico_eq_Ico zero_le_one hp.le, finset.Ico_succ_singleton] }, have hempty : finset.filter p.coprime {0} = ∅, { simp only [finset.filter_singleton, nat.coprime_zero_right, hp.ne', if_false] }, rw [totient_eq_card_coprime, hsplit, finset.filter_union, hempty, finset.empty_union, - ←nat.card_finset_Ico 1 p] at h, + ←nat.card_Ico 1 p] at h, refine p.prime_of_coprime hp (λ n hn hnz, _), apply finset.filter_card_eq h n, refine finset.mem_Ico.mpr ⟨_, hn⟩, diff --git a/src/data/pnat/intervals.lean b/src/data/pnat/intervals.lean index 933196019aaa6..ccfeb6478bb10 100644 --- a/src/data/pnat/intervals.lean +++ b/src/data/pnat/intervals.lean @@ -19,44 +19,45 @@ instance : locally_finite_order ℕ+ := subtype.locally_finite_order _ namespace pnat variables (a b : ℕ+) -@[simp] lemma card_finset_Icc : (Icc a b).card = b + 1 - a := -begin - change (finset.subtype _ _).card = _, - rw [finset.card_subtype, finset.filter_true_of_mem, nat.card_finset_Icc], refl, - exact λ x hx, a.2.trans_le (mem_Icc.1 hx).1, -end - -@[simp] lemma card_finset_Ico : (Ico a b).card = b - a := -begin - change (finset.subtype _ _).card = _, - rw [finset.card_subtype, finset.filter_true_of_mem, nat.card_finset_Ico], refl, - exact λ x hx, a.2.trans_le (mem_Ico.1 hx).1, -end - -@[simp] lemma card_finset_Ioc : (Ioc a b).card = b - a := -begin - change (finset.subtype _ _).card = _, - rw [finset.card_subtype, finset.filter_true_of_mem, nat.card_finset_Ioc], refl, - exact λ x hx, a.2.trans (mem_Ioc.1 hx).1, -end - -@[simp] lemma card_finset_Ioo : (Ioo a b).card = b - a - 1 := -begin - change (finset.subtype _ _).card = _, - rw [finset.card_subtype, finset.filter_true_of_mem, nat.card_finset_Ioo], refl, - exact λ x hx, a.2.trans (mem_Ioo.1 hx).1, -end +lemma Icc_eq_finset_subtype : Icc a b = (Icc (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl +lemma Ico_eq_finset_subtype : Ico a b = (Ico (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl +lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl +lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl + +lemma map_subtype_embedding_Icc : (Icc a b).map (function.embedding.subtype _) = Icc (a : ℕ) b := +map_subtype_embedding_Icc _ _ _ (λ a _ x hx _ ha _, ha.trans_le hx) + +lemma map_subtype_embedding_Ico : (Ico a b).map (function.embedding.subtype _) = Ico (a : ℕ) b := +map_subtype_embedding_Ico _ _ _ (λ a _ x hx _ ha _, ha.trans_le hx) + +lemma map_subtype_embedding_Ioc : (Ioc a b).map (function.embedding.subtype _) = Ioc (a : ℕ) b := +map_subtype_embedding_Ioc _ _ _ (λ a _ x hx _ ha _, ha.trans_le hx) + +lemma map_subtype_embedding_Ioo : (Ioo a b).map (function.embedding.subtype _) = Ioo (a : ℕ) b := +map_subtype_embedding_Ioo _ _ _ (λ a _ x hx _ ha _, ha.trans_le hx) + +@[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := +by rw [←nat.card_Icc, ←map_subtype_embedding_Icc, card_map] + +@[simp] lemma card_Ico : (Ico a b).card = b - a := +by rw [←nat.card_Ico, ←map_subtype_embedding_Ico, card_map] + +@[simp] lemma card_Ioc : (Ioc a b).card = b - a := +by rw [←nat.card_Ioc, ←map_subtype_embedding_Ioc, card_map] + +@[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := +by rw [←nat.card_Ioo, ←map_subtype_embedding_Ioo, card_map] @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := -by rw [←card_finset_Icc, fintype.card_of_finset] +by rw [←card_Icc, fintype.card_of_finset] @[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = b - a := -by rw [←card_finset_Ico, fintype.card_of_finset] +by rw [←card_Ico, fintype.card_of_finset] @[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = b - a := -by rw [←card_finset_Ioc, fintype.card_of_finset] +by rw [←card_Ioc, fintype.card_of_finset] @[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := -by rw [←card_finset_Ioo, fintype.card_of_finset] +by rw [←card_Ioo, fintype.card_of_finset] end pnat diff --git a/src/data/polynomial/iterated_deriv.lean b/src/data/polynomial/iterated_deriv.lean index ee45ff7b96066..cff6f402f0522 100644 --- a/src/data/polynomial/iterated_deriv.lean +++ b/src/data/polynomial/iterated_deriv.lean @@ -123,7 +123,7 @@ lemma coeff_iterated_deriv_as_prod_Ico : ∀ m : ℕ, (iterated_deriv f k).coeff m = (∏ i in Ico m.succ (m + k.succ), i) * (f.coeff (m+k)) := begin induction k with k ih, - { simp only [add_zero, forall_const, one_mul, Ico.self_eq_empty, eq_self_iff_true, + { simp only [add_zero, forall_const, one_mul, Ico_self, eq_self_iff_true, iterated_deriv_zero_right, prod_empty] }, { intro m, rw [iterated_deriv_succ, coeff_derivative, ih (m+1), mul_right_comm], apply congr_arg2, diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index cdcd262dcf4f5..03f971a11df5b 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -25,7 +25,7 @@ In a `locally_finite_order`, * `finset.Ioc`: Open-closed interval as a finset. * `finset.Ioo`: Open-open interval as a finset. * `multiset.Icc`: Closed-closed interval as a multiset. -* `multiset.Ico`: Closed-open interval as a multiset. +* `multiset.Ico`: Closed-open interval as a multiset. Currently only for `ℕ`. * `multiset.Ioc`: Open-closed interval as a multiset. * `multiset.Ioo`: Open-open interval as a finset. @@ -451,21 +451,6 @@ noncomputable def fintype.to_locally_finite_order [fintype α] : finset_mem_Ioc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioc], finset_mem_Ioo := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioo] } -instance [locally_finite_order α] (p : α → Prop) [decidable_pred p] : - locally_finite_order (subtype p) := -{ finset_Icc := λ a b, (Icc a.val b.val).subtype p, - finset_Ico := λ a b, (Ico a.val b.val).subtype p, - finset_Ioc := λ a b, (Ioc a.val b.val).subtype p, - finset_Ioo := λ a b, (Ioo a.val b.val).subtype p, - finset_mem_Icc := λ a b x, by simp_rw [finset.mem_subtype, mem_Icc, subtype.val_eq_coe, - subtype.coe_le_coe], - finset_mem_Ico := λ a b x, by simp_rw [finset.mem_subtype, mem_Ico, subtype.val_eq_coe, - subtype.coe_le_coe, subtype.coe_lt_coe], - finset_mem_Ioc := λ a b x, by simp_rw [finset.mem_subtype, mem_Ioc, subtype.val_eq_coe, - subtype.coe_le_coe, subtype.coe_lt_coe], - finset_mem_Ioo := λ a b x, by simp_rw [finset.mem_subtype, mem_Ioo, subtype.val_eq_coe, - subtype.coe_lt_coe] } - variables [preorder β] [decidable_rel ((≤) : β → β → Prop)] [locally_finite_order β] -- Should this be called `locally_finite_order.lift`? @@ -488,3 +473,68 @@ instance : locally_finite_order (α × β) := finset_mem_Icc := λ a b x, by { rw [mem_product, mem_Icc, mem_Icc, and_and_and_comm], refl } } end preorder + +/-! #### Subtype of a locally finite -/ + +variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] + (p : α → Prop) [decidable_pred p] + +instance : + locally_finite_order (subtype p) := +{ finset_Icc := λ a b, (Icc (a : α) b).subtype p, + finset_Ico := λ a b, (Ico (a : α) b).subtype p, + finset_Ioc := λ a b, (Ioc (a : α) b).subtype p, + finset_Ioo := λ a b, (Ioo (a : α) b).subtype p, + finset_mem_Icc := λ a b x, by simp_rw [finset.mem_subtype, mem_Icc, subtype.coe_le_coe], + finset_mem_Ico := λ a b x, by simp_rw [finset.mem_subtype, mem_Ico, subtype.coe_le_coe, + subtype.coe_lt_coe], + finset_mem_Ioc := λ a b x, by simp_rw [finset.mem_subtype, mem_Ioc, subtype.coe_le_coe, + subtype.coe_lt_coe], + finset_mem_Ioo := λ a b x, by simp_rw [finset.mem_subtype, mem_Ioo, subtype.coe_lt_coe] } + +variables [decidable_eq α] (a b : subtype p) + +namespace finset + +lemma subtype_Icc_eq : Icc a b = (Icc (a : α) b).subtype p := rfl +lemma subtype_Ico_eq : Ico a b = (Ico (a : α) b).subtype p := rfl +lemma subtype_Ioc_eq : Ioc a b = (Ioc (a : α) b).subtype p := rfl +lemma subtype_Ioo_eq : Ioo a b = (Ioo (a : α) b).subtype p := rfl + +variables (hp : ∀ ⦃a b x⦄, a ≤ x → x ≤ b → p a → p b → p x) +include hp + +lemma map_subtype_embedding_Icc : (Icc a b).map (function.embedding.subtype p) = Icc (a : α) b := +begin + rw subtype_Icc_eq, + refine finset.subtype_map_of_mem (λ x hx, _), + rw mem_Icc at hx, + exact hp hx.1 hx.2 a.prop b.prop, +end + +lemma map_subtype_embedding_Ico : (Ico a b).map (function.embedding.subtype p) = Ico (a : α) b := +begin + rw subtype_Ico_eq, + refine finset.subtype_map_of_mem (λ x hx, _), + rw mem_Ico at hx, + exact hp hx.1 hx.2.le a.prop b.prop, +end + +lemma map_subtype_embedding_Ioc : (Ioc a b).map (function.embedding.subtype p) = Ioc (a : α) b := +begin + rw subtype_Ioc_eq, + refine finset.subtype_map_of_mem (λ x hx, _), + rw mem_Ioc at hx, + exact hp hx.1.le hx.2 a.prop b.prop, +end + +lemma map_subtype_embedding_Ioo : (Ioo a b).map (function.embedding.subtype p) = Ioo (a : α) b := +begin + rw subtype_Ioo_eq, + refine finset.subtype_map_of_mem (λ x hx, _), + rw mem_Ioo at hx, + exact hp hx.1.le hx.2.le a.prop b.prop, +end + + +end finset diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index bb8e0c1e89302..07b451fe0989b 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -191,7 +191,7 @@ lemma dist_le_Ico_sum_dist (f : ℕ → α) {m n} (h : m ≤ n) : begin revert n, apply nat.le_induction, - { simp only [finset.sum_empty, finset.Ico.self_eq_empty, dist_self] }, + { simp only [finset.sum_empty, finset.Ico_self, dist_self] }, { assume n hn hrec, calc dist (f m) (f (n+1)) ≤ dist (f m) (f n) + dist _ _ : dist_triangle _ _ _ ... ≤ ∑ i in finset.Ico m n, _ + _ : add_le_add hrec (le_refl _) @@ -202,7 +202,7 @@ end /-- The triangle (polygon) inequality for sequences of points; `finset.range` version. -/ lemma dist_le_range_sum_dist (f : ℕ → α) (n : ℕ) : dist (f 0) (f n) ≤ ∑ i in finset.range n, dist (f i) (f (i + 1)) := -finset.Ico.zero_bot n ▸ dist_le_Ico_sum_dist f (nat.zero_le n) +nat.Ico_zero_eq_range n ▸ dist_le_Ico_sum_dist f (nat.zero_le n) /-- A version of `dist_le_Ico_sum_dist` with each intermediate distance replaced with an upper estimate. -/ @@ -217,7 +217,7 @@ with an upper estimate. -/ lemma dist_le_range_sum_of_dist_le {f : ℕ → α} (n : ℕ) {d : ℕ → ℝ} (hd : ∀ {k}, k < n → dist (f k) (f (k + 1)) ≤ d k) : dist (f 0) (f n) ≤ ∑ i in finset.range n, d i := -finset.Ico.zero_bot n ▸ dist_le_Ico_sum_of_dist_le (zero_le n) (λ _ _, hd) +nat.Ico_zero_eq_range n ▸ dist_le_Ico_sum_of_dist_le (zero_le n) (λ _ _, hd) theorem swap_dist : function.swap (@dist α _) = dist := by funext x y; exact dist_comm _ _ diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index e0cd0e3c00c4d..b66e71cc72d9d 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -130,7 +130,7 @@ lemma edist_le_Ico_sum_edist (f : ℕ → α) {m n} (h : m ≤ n) : begin revert n, refine nat.le_induction _ _, - { simp only [finset.sum_empty, finset.Ico.self_eq_empty, edist_self], + { simp only [finset.sum_empty, finset.Ico_self, edist_self], -- TODO: Why doesn't Lean close this goal automatically? `apply le_refl` fails too. exact le_refl (0:ℝ≥0∞) }, { assume n hn hrec, @@ -143,7 +143,7 @@ end /-- The triangle (polygon) inequality for sequences of points; `finset.range` version. -/ lemma edist_le_range_sum_edist (f : ℕ → α) (n : ℕ) : edist (f 0) (f n) ≤ ∑ i in finset.range n, edist (f i) (f (i + 1)) := -finset.Ico.zero_bot n ▸ edist_le_Ico_sum_edist f (nat.zero_le n) +nat.Ico_zero_eq_range n ▸ edist_le_Ico_sum_edist f (nat.zero_le n) /-- A version of `edist_le_Ico_sum_edist` with each intermediate distance replaced with an upper estimate. -/ @@ -158,7 +158,7 @@ with an upper estimate. -/ lemma edist_le_range_sum_of_edist_le {f : ℕ → α} (n : ℕ) {d : ℕ → ℝ≥0∞} (hd : ∀ {k}, k < n → edist (f k) (f (k + 1)) ≤ d k) : edist (f 0) (f n) ≤ ∑ i in finset.range n, d i := -finset.Ico.zero_bot n ▸ edist_le_Ico_sum_of_edist_le (zero_le n) (λ _ _, hd) +nat.Ico_zero_eq_range n ▸ edist_le_Ico_sum_of_edist_le (zero_le n) (λ _ _, hd) /-- Reformulation of the uniform structure in terms of the extended distance -/ theorem uniformity_pseudoedist : From 37695864c2508bf1682e063414cd8b5ceb9a32a2 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Fri, 17 Sep 2021 14:45:46 +0200 Subject: [PATCH 12/41] fix ADE_inequality --- src/number_theory/ADE_inequality.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/number_theory/ADE_inequality.lean b/src/number_theory/ADE_inequality.lean index 13906a3a352d9..20913ffcd0fbe 100644 --- a/src/number_theory/ADE_inequality.lean +++ b/src/number_theory/ADE_inequality.lean @@ -5,7 +5,7 @@ Authors: Johan Commelin -/ import data.multiset.sort -import data.pnat.basic +import data.pnat.intervals import data.rat.order import tactic.norm_num From 91a6ccafb911f0bb4edfaf25092393a7a9e7a2e5 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Fri, 17 Sep 2021 15:25:17 +0200 Subject: [PATCH 13/41] fix long line --- src/data/nat/intervals.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/data/nat/intervals.lean b/src/data/nat/intervals.lean index 18c9e7c14f6d0..a452456288df3 100644 --- a/src/data/nat/intervals.lean +++ b/src/data/nat/intervals.lean @@ -142,7 +142,8 @@ begin split, { rintro ⟨x, hx, rfl⟩, rw mem_Ico at hx, - refine ⟨ _, ((nat.sub_le_sub_left_iff hac).2 hx.1).trans_lt ((nat.sub_lt_sub_right_iff hac).2 (nat.lt_succ_self _))⟩, + refine ⟨_, ((nat.sub_le_sub_left_iff hac).2 hx.1).trans_lt ((nat.sub_lt_sub_right_iff hac).2 + (nat.lt_succ_self _))⟩, cases lt_or_le c b, { rw nat.sub_eq_zero_of_le h, exact zero_le _ }, From b486e33d831b5c5df876a19d952089b41e0c9152 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 17 Sep 2021 10:39:12 -0400 Subject: [PATCH 14/41] Replace fancy default fields with a custom constructor Also drop `decidable_rel` assumption to avoid diamonds --- src/order/locally_finite.lean | 123 ++++++++++++---------------------- 1 file changed, 42 insertions(+), 81 deletions(-) diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 03f971a11df5b..cc5c98cdd4e35 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -101,65 +101,29 @@ situation is more complicated, as we need to call an hypothesis (namely our use of `id_annotate` to force it to appear there. -/ -section tactic -open tactic - -variables {α : Type*} [preorder α] [decidable_rel ((≤) : α → α → Prop)] - {finset_Icc : α → α → finset α} (finset_mem_Icc : ∀ a b x : α, x ∈ finset_Icc a b ↔ a ≤ x ∧ x ≤ b) - -/-- Auxiliary lemma to provide `locally_finite_order.mem_finset_Ico` when -`locally_finite_order.finset_Ico` is the default value. -/ -lemma locally_finite_order.mem_Ico_of_mem_Icc : - id_annotate finset_mem_Icc - (∀ a b x : α, x ∈ (finset_Icc a b).filter (λ x, ¬b ≤ x) ↔ a ≤ x ∧ x < b) := -λ a b x, by rw [finset.mem_filter, finset_mem_Icc, and_assoc, lt_iff_le_not_le] - -/-- Auxiliary lemma to provide `locally_finite_order.mem_finset_Ioc` when -`locally_finite_order.finset_Ioc` is the default value. -/ -lemma locally_finite_order.mem_Ioc_of_mem_Icc : - id_annotate finset_mem_Icc - (∀ a b x : α, x ∈ (finset_Icc a b).filter (λ x, ¬x ≤ a) ↔ a < x ∧ x ≤ b) := -λ a b x, by rw [finset.mem_filter, finset_mem_Icc, and.right_comm, lt_iff_le_not_le] - -/-- Auxiliary lemma to provide `locally_finite_order.mem_finset_Ioo` when -`locally_finite_order.finset_Ico` is the default value. -/ -lemma locally_finite_order.mem_Ioo_of_mem_Icc : - id_annotate finset_mem_Icc - (∀ a b x : α, x ∈ (finset_Icc a b).filter (λ x, ¬x ≤ a ∧ ¬b ≤ x) ↔ a < x ∧ x < b) := -λ a b x, by rw [finset.mem_filter, finset_mem_Icc, and_and_and_comm, lt_iff_le_not_le, - lt_iff_le_not_le] - -/-- Here to make `locally_finite_order.finset_Ico` an optional field. Tries to prove -`locally_finite_order.mem_finset_Ico` -/ -meta def finset_Ico_laws_tac : tactic unit := -`[exact locally_finite_order.mem_Ico_of_mem_Icc _] - -/-- Here to make `locally_finite_order.finset_Ioc` an optional field. Tries to prove -`locally_finite_order.mem_finset_Ioc` -/ -meta def finset_Ioc_laws_tac : tactic unit := -`[exact locally_finite_order.mem_Ioc_of_mem_Icc _] - -/-- Here to make `locally_finite_order.finset_Ioo` an optional field. Tries to prove -`locally_finite_order.mem_finset_Ioo` -/ -meta def finset_Ioo_laws_tac : tactic unit := -`[exact locally_finite_order.mem_Ioo_of_mem_Icc _] - -end tactic - open finset -class locally_finite_order (α : Type*) [preorder α] [decidable_rel ((≤) : α → α → Prop)] := +class locally_finite_order (α : Type*) [preorder α] := (finset_Icc : α → α → finset α) -(finset_Ico : α → α → finset α := λ a b, (finset_Icc a b).filter (λ x, ¬b ≤ x)) -(finset_Ioc : α → α → finset α := λ a b, (finset_Icc a b).filter (λ x, ¬x ≤ a)) -(finset_Ioo : α → α → finset α := λ a b, (finset_Icc a b).filter (λ x, ¬x ≤ a ∧ ¬b ≤ x)) +(finset_Ico : α → α → finset α) +(finset_Ioc : α → α → finset α) +(finset_Ioo : α → α → finset α) (finset_mem_Icc : ∀ a b x : α, x ∈ finset_Icc a b ↔ a ≤ x ∧ x ≤ b) -(finset_mem_Ico : id_annotate finset_mem_Icc (∀ a b x : α, x ∈ finset_Ico a b ↔ a ≤ x ∧ x < b) - . finset_Ico_laws_tac) -(finset_mem_Ioc : id_annotate finset_mem_Icc (∀ a b x : α, x ∈ finset_Ioc a b ↔ a < x ∧ x ≤ b) - . finset_Ioc_laws_tac) -(finset_mem_Ioo : id_annotate finset_mem_Icc (∀ a b x : α, x ∈ finset_Ioo a b ↔ a < x ∧ x < b) - . finset_Ioo_laws_tac) +(finset_mem_Ico : ∀ a b x : α, x ∈ finset_Ico a b ↔ a ≤ x ∧ x < b) +(finset_mem_Ioc : ∀ a b x : α, x ∈ finset_Ioc a b ↔ a < x ∧ x ≤ b) +(finset_mem_Ioo : ∀ a b x : α, x ∈ finset_Ioo a b ↔ a < x ∧ x < b) + +def locally_finite_order.of_Icc (α : Type*) [preorder α] [decidable_rel ((≤) : α → α → Prop)] + (finset_Icc : α → α → finset α) (mem_Icc : ∀ a b x, x ∈ finset_Icc a b ↔ a ≤ x ∧ x ≤ b) : + locally_finite_order α := +{ finset_Icc := finset_Icc, + finset_Ico := λ a b, (finset_Icc a b).filter (λ x, ¬b ≤ x), + finset_Ioc := λ a b, (finset_Icc a b).filter (λ x, ¬x ≤ a), + finset_Ioo := λ a b, (finset_Icc a b).filter (λ x, ¬x ≤ a ∧ ¬b ≤ x), + finset_mem_Icc := mem_Icc, + finset_mem_Ico := λ a b x, by simp [mem_Icc, lt_iff_le_not_le, and_assoc], + finset_mem_Ioc := λ a b x, by simp [mem_Icc, lt_iff_le_not_le, and.right_comm], + finset_mem_Ioo := λ a b x, by { simp only [mem_Icc, lt_iff_le_not_le, mem_filter], tauto } } variables {α β : Type*} @@ -167,7 +131,7 @@ variables {α β : Type*} namespace finset section preorder -variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] +variables [preorder α] [locally_finite_order α] /-- The finset of elements `x` such that `a ≤ x` and `x ≤ b`. Basically `set.Icc a b` as a finset. -/ @@ -224,21 +188,18 @@ end end preorder section partial_order -variables [partial_order α] [decidable_rel ((≤) : α → α → Prop)] - [locally_finite_order α] +variables [partial_order α] [locally_finite_order α] end partial_order section order_top -variables [order_top α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] +variables [order_top α] [locally_finite_order α] /-- The finset of elements `x` such that `a ≤ x`. Basically `set.Ici a` as a finset. -/ -def Ici (a : α) : finset α := -Icc a ⊤ +def Ici (a : α) : finset α := Icc a ⊤ /-- The finset of elements `x` such that `a < x`. Basically `set.Ioi a` as a finset. -/ -def Ioi (a : α) : finset α := -Ioc a ⊤ +def Ioi (a : α) : finset α := Ioc a ⊤ lemma Ici_eq_Icc (a : α) : Ici a = Icc a ⊤ := rfl lemma Ioi_eq_Ioc (a : α) : Ioi a = Ioc a ⊤ := rfl @@ -258,7 +219,7 @@ by rw [←set.mem_Ioi, ←coe_Ioi, mem_coe] end order_top section order_bot -variables [order_bot α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] +variables [order_bot α] [locally_finite_order α] /-- The finset of elements `x` such that `x ≤ b`. Basically `set.Iic b` as a finset. -/ def Iic (b : α) : finset α := @@ -290,7 +251,7 @@ end finset namespace multiset section preorder -variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] +variables [preorder α] [locally_finite_order α] /-- The multiset of elements `x` such that `a ≤ x` and `x ≤ b`. Basically `set.Icc a b` as a multiset. -/ @@ -321,7 +282,7 @@ by rw [Ioo, ←finset.mem_def, finset.mem_Ioo] end preorder section order_top -variables [order_top α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] +variables [order_top α] [locally_finite_order α] /-- The multiset of elements `x` such that `a ≤ x`. Basically `set.Ici a` as a multiset. -/ def Ici (a : α) : multiset α := @@ -340,7 +301,7 @@ by rw [Ioi, ←finset.mem_def, finset.mem_Ioi] end order_top section order_bot -variables [order_bot α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] +variables [order_bot α] [locally_finite_order α] /-- The multiset of elements `x` such that `x ≤ b`. Basically `set.Iic b` as a multiset. -/ def Iic (b : α) : multiset α := @@ -363,7 +324,7 @@ end multiset namespace set section preorder -variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] (a b : α) +variables [preorder α] [locally_finite_order α] (a b : α) instance fintype_Icc : fintype (Icc a b) := fintype.of_finset (finset.Icc a b) (λ x, by rw [finset.mem_Icc, mem_Icc]) @@ -392,7 +353,7 @@ lemma finite_Ioo : (Ioo a b).finite := end preorder section order_top -variables [order_top α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] (a : α) +variables [order_top α] [locally_finite_order α] (a : α) instance fintype_Ici : fintype (Ici a) := fintype.of_finset (finset.Ici a) (λ x, by rw [finset.mem_Ici, mem_Ici]) @@ -409,7 +370,7 @@ lemma finite_Ioi : (Ioi a).finite := end order_top section order_bot -variables [order_bot α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] (b : α) +variables [order_bot α] [locally_finite_order α] (b : α) instance fintype_Iic : fintype (Iic b) := fintype.of_finset (finset.Iic b) (λ x, by rw [finset.mem_Iic, mem_Iic]) @@ -432,13 +393,14 @@ end set open finset section preorder -variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] +variables [preorder α] noncomputable def locally_finite_order_of_finite_Icc (h : ∀ a b : α, (set.Icc a b).finite) : locally_finite_order α := -{ finset_Icc := λ a b, (h a b).to_finset, - finset_mem_Icc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Icc] } +@locally_finite_order.of_Icc α _ (classical.dec_rel _) + (λ a b, (h a b).to_finset) + (λ a b x, by rw [set.finite.mem_to_finset, set.mem_Icc]) noncomputable def fintype.to_locally_finite_order [fintype α] : locally_finite_order α := @@ -451,7 +413,7 @@ noncomputable def fintype.to_locally_finite_order [fintype α] : finset_mem_Ioc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioc], finset_mem_Ioo := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioo] } -variables [preorder β] [decidable_rel ((≤) : β → β → Prop)] [locally_finite_order β] +variables [preorder β] [locally_finite_order β] -- Should this be called `locally_finite_order.lift`? /-- Given an order embedding `α ↪o β`, pulls back the `locally_finite_order` on `β` to `α`. -/ @@ -468,19 +430,18 @@ noncomputable def order_embedding.locally_finite_order (f : α ↪o β) : variables [locally_finite_order α] -instance : locally_finite_order (α × β) := -{ finset_Icc := λ a b, (Icc a.fst b.fst).product (Icc a.snd b.snd), - finset_mem_Icc := λ a b x, by { rw [mem_product, mem_Icc, mem_Icc, and_and_and_comm], refl } } +instance [decidable_rel ((≤) : α × β → α × β → Prop)] : locally_finite_order (α × β) := +locally_finite_order.of_Icc (α × β) + (λ a b, (Icc a.fst b.fst).product (Icc a.snd b.snd)) + (λ a b x, by { rw [mem_product, mem_Icc, mem_Icc, and_and_and_comm], refl }) end preorder /-! #### Subtype of a locally finite -/ -variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] - (p : α → Prop) [decidable_pred p] +variables [preorder α] [locally_finite_order α] (p : α → Prop) [decidable_pred p] -instance : - locally_finite_order (subtype p) := +instance : locally_finite_order (subtype p) := { finset_Icc := λ a b, (Icc (a : α) b).subtype p, finset_Ico := λ a b, (Ico (a : α) b).subtype p, finset_Ioc := λ a b, (Ioc (a : α) b).subtype p, From 176394c24fd553353a46768804c53e75ec2b3bb4 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Fri, 17 Sep 2021 18:37:05 +0200 Subject: [PATCH 15/41] polish docstrings --- src/data/fin/intervals.lean | 65 ++++++++++++++++++++++++++ src/data/finset/intervals.lean | 33 ++++++------- src/data/int/intervals.lean | 2 +- src/data/nat/intervals.lean | 2 +- src/data/pnat/intervals.lean | 10 ++-- src/order/locally_finite.lean | 84 ++++++++++++++++++---------------- 6 files changed, 132 insertions(+), 64 deletions(-) create mode 100644 src/data/fin/intervals.lean diff --git a/src/data/fin/intervals.lean b/src/data/fin/intervals.lean new file mode 100644 index 0000000000000..650ec889d18da --- /dev/null +++ b/src/data/fin/intervals.lean @@ -0,0 +1,65 @@ +/- +Copyright (c) 2020 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison, Yaël Dillies +-/ +import data.nat.intervals + +/-! +# Finite intervals in `fin n` + +This file proves that `fin n` is a `locally_finite_order` and calculates the cardinality of its +intervals as finsets and fintypes. +-/ + +open finset fin + +variables (n : ℕ) + +instance : locally_finite_order (fin n) := subtype.locally_finite_order _ + +namespace fin +variables {n} (a b : fin n) + +lemma Icc_eq_finset_subtype : Icc a b = (Icc (a : ℕ) b).subtype (λ x, x < n) := rfl +lemma Ico_eq_finset_subtype : Ico a b = (Ico (a : ℕ) b).subtype (λ x, x < n) := rfl +lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).subtype (λ x, x < n) := rfl +lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).subtype (λ x, x < n) := rfl + +lemma map_subtype_embedding_Icc : (Icc a b).map (function.embedding.subtype _) = Icc (a : ℕ) b := +map_subtype_embedding_Icc _ _ _ (λ _ c x _ hx _, hx.trans_lt) + +lemma map_subtype_embedding_Ico : (Ico a b).map (function.embedding.subtype _) = Ico (a : ℕ) b := +map_subtype_embedding_Ico _ _ _ (λ _ c x _ hx _, hx.trans_lt) + +lemma map_subtype_embedding_Ioc : (Ioc a b).map (function.embedding.subtype _) = Ioc (a : ℕ) b := +map_subtype_embedding_Ioc _ _ _ (λ _ c x _ hx _, hx.trans_lt) + +lemma map_subtype_embedding_Ioo : (Ioo a b).map (function.embedding.subtype _) = Ioo (a : ℕ) b := +map_subtype_embedding_Ioo _ _ _ (λ _ c x _ hx _, hx.trans_lt) + +@[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := +by rw [←nat.card_Icc, ←map_subtype_embedding_Icc, card_map] + +@[simp] lemma card_Ico : (Ico a b).card = b - a := +by rw [←nat.card_Ico, ←map_subtype_embedding_Ico, card_map] + +@[simp] lemma card_Ioc : (Ioc a b).card = b - a := +by rw [←nat.card_Ioc, ←map_subtype_embedding_Ioc, card_map] + +@[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := +by rw [←nat.card_Ioo, ←map_subtype_embedding_Ioo, card_map] + +@[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := +by rw [←card_Icc, fintype.card_of_finset] + +@[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = b - a := +by rw [←card_Ico, fintype.card_of_finset] + +@[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = b - a := +by rw [←card_Ioc, fintype.card_of_finset] + +@[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := +by rw [←card_Ioo, fintype.card_of_finset] + +end fin diff --git a/src/data/finset/intervals.lean b/src/data/finset/intervals.lean index 60fbf36b4a340..eb4d4db969454 100644 --- a/src/data/finset/intervals.lean +++ b/src/data/finset/intervals.lean @@ -3,26 +3,21 @@ Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Yaël Dillies -/ -import algebra.ordered_sub -import data.finset.basic -import data.multiset.intervals import order.locally_finite /-! # Intervals as finsets -For now this only covers `Ico a b`, the "closed-open" interval containing `[a, ..., b - 1]`. +This file provides basic results about all the `finset.Ixx`, which are defined in +`order.locally_finite`. ## TODO -This file was originally only about `finset.Ico a b` where `a n : ℕ`. No care has yet been taken to -generalize these lemmas properly and many lemmas about `Icc`, `Ioc`, `Ioo` are missing. +This file was originally only about `finset.Ico a b` where `a b : ℕ`. No care has yet been taken to +generalize these lemmas properly and many lemmas about `Icc`, `Ioc`, `Ioo` are missing. In general, +what's to do is taking the lemmas in `data.x.intervals` and abstract away the concrete structure. -/ --- TODO: There is a diamond between `nat.decidable_eq` and `decidable_eq_of_decidable_le` --- and another one between `linear_order.decidable_lt` and `decidable_lt_of_decidable_le` -local attribute [-instance] decidable_eq_of_decidable_le decidable_lt_of_decidable_le - lemma finset.coe_eq_singleton {α : Type*} {s : finset α} {a : α} : (s : set α) = {a} ↔ s = {a} := by rw [←finset.coe_singleton, finset.coe_inj] @@ -30,7 +25,7 @@ by rw [←finset.coe_singleton, finset.coe_inj] namespace finset variables {α : Type*} section preorder -variables [preorder α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] {a b : α} +variables [preorder α] [locally_finite_order α] {a b : α} @[simp] lemma nonempty_Icc : (Icc a b).nonempty ↔ a ≤ b := by rw [←coe_nonempty, coe_Icc, set.nonempty_Icc] @@ -75,7 +70,7 @@ Ioc_eq_empty h.not_lt @[simp] lemma Ioo_eq_empty_of_le (h : b ≤ a) : Ioo a b = ∅ := Ioo_eq_empty h.not_lt -variables (a b) +variables (a) @[simp] lemma Ico_self : Ico a a = ∅ := by rw [←coe_eq_empty, coe_Ico, set.Ico_self] @@ -105,14 +100,15 @@ begin exact and_iff_left_of_imp (λ h, h.2.trans_le hcb), end -lemma Ico_filter_le_of_le_left {a b c : α} (hca : c ≤ a) : +lemma Ico_filter_le_of_le_left [decidable_rel ((≤) : α → α → Prop)] {a b c : α} (hca : c ≤ a) : (Ico a b).filter (λ x, c ≤ x) = Ico a b := finset.filter_true_of_mem (λ x hx, hca.trans (mem_Ico.1 hx).1) -lemma Ico_filter_le_of_right_le {a b c : α} : (Ico a b).filter (λ x, b ≤ x) = ∅ := +lemma Ico_filter_le_of_right_le [decidable_rel ((≤) : α → α → Prop)] {a b : α} : + (Ico a b).filter (λ x, b ≤ x) = ∅ := finset.filter_false_of_mem (λ x hx, (mem_Ico.1 hx).2.not_le) -lemma Ico_filter_le_of_left_le {a b c : α} (hac : a ≤ c) : +lemma Ico_filter_le_of_left_le [decidable_rel ((≤) : α → α → Prop)] {a b c : α} (hac : a ≤ c) : (Ico a b).filter (λ x, c ≤ x) = Ico c b := begin ext x, @@ -123,7 +119,7 @@ end end preorder section partial_order -variables [partial_order α] [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] {a b : α} +variables [partial_order α] [locally_finite_order α] {a b : α} @[simp] lemma Icc_self (a : α) : Icc a a = {a} := @@ -145,7 +141,8 @@ end lemma Ico_disjoint_Ico_consecutive [decidable_eq α] (a b c : α) : disjoint (Ico a b) (Ico b c) := le_of_eq $ Ico_inter_Ico_consecutive a b c -lemma Ico_filter_le_left {a b : α} (hab : a < b) : (Ico a b).filter (λ x, x ≤ a) = {a} := +lemma Ico_filter_le_left [decidable_rel ((≤) : α → α → Prop)] {a b : α} (hab : a < b) : + (Ico a b).filter (λ x, x ≤ a) = {a} := begin ext x, rw [mem_filter, mem_Ico, mem_singleton, and.right_comm, ←le_antisymm_iff, eq_comm], @@ -212,7 +209,7 @@ end linear_order section ordered_cancel_add_comm_monoid variables [ordered_cancel_add_comm_monoid α] [has_exists_add_of_le α] [decidable_eq α] - [decidable_rel ((≤) : α → α → Prop)] [locally_finite_order α] + [locally_finite_order α] lemma image_add_const_Icc (a b c : α) : (Icc a b).image ((+) c) = Icc (a + c) (b + c) := begin diff --git a/src/data/int/intervals.lean b/src/data/int/intervals.lean index 52e0f14c1b62e..13bb7a54862af 100644 --- a/src/data/int/intervals.lean +++ b/src/data/int/intervals.lean @@ -7,7 +7,7 @@ import data.int.basic import data.nat.intervals /-! -# Intervals of integers +# Finite intervals of integers This file proves that `ℤ` is a `locally_finite_order` and calculates the cardinality of its intervals as finsets and fintypes. diff --git a/src/data/nat/intervals.lean b/src/data/nat/intervals.lean index a452456288df3..2f13479164638 100644 --- a/src/data/nat/intervals.lean +++ b/src/data/nat/intervals.lean @@ -6,7 +6,7 @@ Authors: Yaël Dillies import data.finset.intervals /-! -# Intervals of naturals +# Finite intervals of naturals This file proves that `ℕ` is a `locally_finite_order` and calculates the cardinality of its intervals as finsets and fintypes. diff --git a/src/data/pnat/intervals.lean b/src/data/pnat/intervals.lean index ccfeb6478bb10..ca6923a32d374 100644 --- a/src/data/pnat/intervals.lean +++ b/src/data/pnat/intervals.lean @@ -6,7 +6,7 @@ Authors: Scott Morrison, Yaël Dillies import data.nat.intervals /-! -# Intervals of positive naturals +# Finite intervals of positive naturals This file proves that `ℕ+` is a `locally_finite_order` and calculates the cardinality of its intervals as finsets and fintypes. @@ -25,16 +25,16 @@ lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).subtype (λ (n : ℕ), lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl lemma map_subtype_embedding_Icc : (Icc a b).map (function.embedding.subtype _) = Icc (a : ℕ) b := -map_subtype_embedding_Icc _ _ _ (λ a _ x hx _ ha _, ha.trans_le hx) +map_subtype_embedding_Icc _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) lemma map_subtype_embedding_Ico : (Ico a b).map (function.embedding.subtype _) = Ico (a : ℕ) b := -map_subtype_embedding_Ico _ _ _ (λ a _ x hx _ ha _, ha.trans_le hx) +map_subtype_embedding_Ico _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) lemma map_subtype_embedding_Ioc : (Ioc a b).map (function.embedding.subtype _) = Ioc (a : ℕ) b := -map_subtype_embedding_Ioc _ _ _ (λ a _ x hx _ ha _, ha.trans_le hx) +map_subtype_embedding_Ioc _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) lemma map_subtype_embedding_Ioo : (Ioo a b).map (function.embedding.subtype _) = Ioo (a : ℕ) b := -map_subtype_embedding_Ioo _ _ _ (λ a _ x hx _ ha _, ha.trans_le hx) +map_subtype_embedding_Ioo _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) @[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := by rw [←nat.card_Icc, ←map_subtype_embedding_Icc, card_map] diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index cc5c98cdd4e35..0be8d01f55cca 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -4,8 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import data.finset.preimage -import data.set.finite -import data.set.intervals.basic /-! # Locally finite orders @@ -43,18 +41,25 @@ When it's also an `order_bot`, ## Instances -We provide a `locally_finite_order` instance for -* a subtype of a locally finite order -* any fintype (but it is noncomputable), see `fintype.locally_finite_order` +A `locally_finite_order` instance can be built +* for a subtype of a locally finite order. See `subtype.locally_finite_order`. +* for the product of two locally finite orders. See `prod.locally_finite_order`. +* for any fintype (but it is noncomputable). See `fintype.locally_finite_order`. +* from a definition of `finset.Icc` alone. See `locally_finite_order.of_Icc`. +* by pulling back `locally_finite_order β` through an order embedding `f : α →o β`. See + `order_embedding.locally_finite_order`. Instances for concrete types are proved in their respective files: -* `data.nat.intervals` for `ℕ` -* `data.int.intervals` for `ℤ` -* `data.pnat.intervals` for `ℕ+` +* `ℕ` is in `data.nat.intervals` +* `ℤ` is in `data.int.intervals` +* `ℕ+` is in `data.pnat.intervals` +* `fin n` is in `data.fin.intervals` Along, you will find lemmas about the cardinality of those finite intervals. ## TODO +`multiset.Ico` hasn't been generalized yet. All of `data.multiset.intervals` should be generalized. + Provide the `locally_finite_order` instance for `lex α β`. From `linear_order α`, `no_top_order α`, `locally_finite_order α`, we can also define an @@ -85,24 +90,11 @@ successor (and actually a predecessor as well), so it is a `succ_order`, but it' as `Icc (-1) 1` is infinite. -/ -/-! -### Arguments autofilling tactics - -`locally_finite_order` has fields for all of `Icc`, `Ico`, `Ioc`, `Ioo` to let control over -definitional equality. But all fields can be recovered from `Icc` (or any other actually, but we had -to make a choice here). Hence we provide default values for `locally_finite_order.finset_Ico`, -`locally_finite_order.finset_Ioc`, `locally_finite_order.finset_Ioo` and three corresponding tactics -to fill in the three corresponding proofs. - -This is analogous to the default definition of `(<)` in terms of `(≤)`. There, the default proof of -`lt_iff_le_not_le` is provided by `order_laws_tac`, which just tries to call `refl`. Here, the -situation is more complicated, as we need to call an hypothesis (namely -`locally_finite_order.mem_finset_Icc`) that doesn't appear a priori in the type of the proof. Whence -our use of `id_annotate` to force it to appear there. --/ - open finset +/-- A locally finite order is an order where bounded intervals are finite. When you don't care too +much about definitional equality, you can use `locally_finite_order.of_Icc` or +`locally_finite_order.of_finite_Icc` to build a locally finite order from just `finset.Icc`. -/ class locally_finite_order (α : Type*) [preorder α] := (finset_Icc : α → α → finset α) (finset_Ico : α → α → finset α) @@ -113,7 +105,9 @@ class locally_finite_order (α : Type*) [preorder α] := (finset_mem_Ioc : ∀ a b x : α, x ∈ finset_Ioc a b ↔ a < x ∧ x ≤ b) (finset_mem_Ioo : ∀ a b x : α, x ∈ finset_Ioo a b ↔ a < x ∧ x < b) -def locally_finite_order.of_Icc (α : Type*) [preorder α] [decidable_rel ((≤) : α → α → Prop)] +/-- A constructor from a definition of `finset.Icc` alone, the other ones being derived by removing +the ends. This one requires `decidable_rel (≤)` but only `preorder`. -/ +def locally_finite_order.of_Icc' (α : Type*) [preorder α] [decidable_rel ((≤) : α → α → Prop)] (finset_Icc : α → α → finset α) (mem_Icc : ∀ a b x, x ∈ finset_Icc a b ↔ a ≤ x ∧ x ≤ b) : locally_finite_order α := { finset_Icc := finset_Icc, @@ -121,9 +115,25 @@ def locally_finite_order.of_Icc (α : Type*) [preorder α] [decidable_rel ((≤) finset_Ioc := λ a b, (finset_Icc a b).filter (λ x, ¬x ≤ a), finset_Ioo := λ a b, (finset_Icc a b).filter (λ x, ¬x ≤ a ∧ ¬b ≤ x), finset_mem_Icc := mem_Icc, - finset_mem_Ico := λ a b x, by simp [mem_Icc, lt_iff_le_not_le, and_assoc], - finset_mem_Ioc := λ a b x, by simp [mem_Icc, lt_iff_le_not_le, and.right_comm], - finset_mem_Ioo := λ a b x, by { simp only [mem_Icc, lt_iff_le_not_le, mem_filter], tauto } } + finset_mem_Ico := λ a b x, by rw [finset.mem_filter, mem_Icc, and_assoc, lt_iff_le_not_le], + finset_mem_Ioc := λ a b x, by rw [finset.mem_filter, mem_Icc, and.right_comm, lt_iff_le_not_le], + finset_mem_Ioo := λ a b x, by rw [finset.mem_filter, mem_Icc, and_and_and_comm, lt_iff_le_not_le, + lt_iff_le_not_le] } + +/-- A constructor from a definition of `finset.Icc` alone, the other ones being derived by removing +the ends. This one requires `partial_order` but only `decidable_eq`. -/ +def locally_finite_order.of_Icc (α : Type*) [partial_order α] [decidable_eq α] + (finset_Icc : α → α → finset α) (mem_Icc : ∀ a b x, x ∈ finset_Icc a b ↔ a ≤ x ∧ x ≤ b) : + locally_finite_order α := +{ finset_Icc := finset_Icc, + finset_Ico := λ a b, (finset_Icc a b).filter (λ x, x ≠ b), + finset_Ioc := λ a b, (finset_Icc a b).filter (λ x, a ≠ x), + finset_Ioo := λ a b, (finset_Icc a b).filter (λ x, a ≠ x ∧ x ≠ b), + finset_mem_Icc := mem_Icc, + finset_mem_Ico := λ a b x, by rw [finset.mem_filter, mem_Icc, and_assoc, lt_iff_le_and_ne], + finset_mem_Ioc := λ a b x, by rw [finset.mem_filter, mem_Icc, and.right_comm, lt_iff_le_and_ne], + finset_mem_Ioo := λ a b x, by rw [finset.mem_filter, mem_Icc, and_and_and_comm, lt_iff_le_and_ne, + lt_iff_le_and_ne] } variables {α β : Type*} @@ -187,11 +197,6 @@ end end preorder -section partial_order -variables [partial_order α] [locally_finite_order α] - -end partial_order - section order_top variables [order_top α] [locally_finite_order α] @@ -395,13 +400,15 @@ open finset section preorder variables [preorder α] -noncomputable def locally_finite_order_of_finite_Icc +/-- A noncomputable constructor from the finiteness of all closed intervals. -/ +noncomputable def locally_finite_order.of_finite_Icc (h : ∀ a b : α, (set.Icc a b).finite) : locally_finite_order α := -@locally_finite_order.of_Icc α _ (classical.dec_rel _) +@locally_finite_order.of_Icc' α _ (classical.dec_rel _) (λ a b, (h a b).to_finset) (λ a b x, by rw [set.finite.mem_to_finset, set.mem_Icc]) +/-- A fintype is noncomputably a locally finite order. -/ noncomputable def fintype.to_locally_finite_order [fintype α] : locally_finite_order α := { finset_Icc := λ a b, (set.finite.of_fintype (set.Icc a b)).to_finset, @@ -431,13 +438,13 @@ noncomputable def order_embedding.locally_finite_order (f : α ↪o β) : variables [locally_finite_order α] instance [decidable_rel ((≤) : α × β → α × β → Prop)] : locally_finite_order (α × β) := -locally_finite_order.of_Icc (α × β) +locally_finite_order.of_Icc' (α × β) (λ a b, (Icc a.fst b.fst).product (Icc a.snd b.snd)) (λ a b x, by { rw [mem_product, mem_Icc, mem_Icc, and_and_and_comm], refl }) end preorder -/-! #### Subtype of a locally finite -/ +/-! #### Subtype of a locally finite order -/ variables [preorder α] [locally_finite_order α] (p : α → Prop) [decidable_pred p] @@ -453,7 +460,7 @@ instance : locally_finite_order (subtype p) := subtype.coe_lt_coe], finset_mem_Ioo := λ a b x, by simp_rw [finset.mem_subtype, mem_Ioo, subtype.coe_lt_coe] } -variables [decidable_eq α] (a b : subtype p) +variables (a b : subtype p) namespace finset @@ -497,5 +504,4 @@ begin exact hp hx.1.le hx.2.le a.prop b.prop, end - end finset From 76583648ae6b2a9a530b733d5c81ee4f293fb420 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Fri, 17 Sep 2021 23:23:21 +0200 Subject: [PATCH 16/41] fix divisors --- src/number_theory/divisors.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/src/number_theory/divisors.lean b/src/number_theory/divisors.lean index 35e0556c6fa81..807eab2902083 100644 --- a/src/number_theory/divisors.lean +++ b/src/number_theory/divisors.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import algebra.big_operators.order -import data.finset.intervals +import data.nat.intervals import data.nat.prime /-! @@ -70,7 +70,8 @@ end lemma divisors_eq_proper_divisors_insert_self_of_pos (h : 0 < n): divisors n = has_insert.insert n (proper_divisors n) := -by rw [divisors, proper_divisors, finset.Ico_insert_right h, finset.filter_insert, if_pos (dvd_refl n)] +by rw [divisors, proper_divisors, Ico_succ_right_eq_insert_Ico h, finset.filter_insert, + if_pos (dvd_refl n)] @[simp] lemma mem_divisors {m : ℕ} : From 59c5e90c3219fb2a57ace4d796e973c136d64e81 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sat, 18 Sep 2021 08:20:30 +0200 Subject: [PATCH 17/41] fix primorial --- src/number_theory/primorial.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/number_theory/primorial.lean b/src/number_theory/primorial.lean index 2a3e08dd573ca..0a8783bf19990 100644 --- a/src/number_theory/primorial.lean +++ b/src/number_theory/primorial.lean @@ -115,8 +115,7 @@ lemma primorial_le_4_pow : ∀ (n : ℕ), n# ≤ 4 ^ n = ∏ i in filter prime (range (2 * m + 2)), i : by simpa [←twice_m] ... = ∏ i in filter prime (finset.Ico (m + 2) (2 * m + 2) ∪ range (m + 2)), i : begin - rw [range_eq_Ico, range_eq_Ico, finset.union_comm, - finset.Ico_union_Ico_consecutive], + rw [range_eq_Ico, finset.union_comm, finset.Ico_union_Ico_eq_Ico], exact bot_le, simp only [add_le_add_iff_right], linarith, From 62bce1575e80d82bfc73f53209f90c3d375e00ab Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sat, 18 Sep 2021 11:16:39 +0200 Subject: [PATCH 18/41] fix imports --- src/data/polynomial/inductions.lean | 2 +- src/data/polynomial/iterated_deriv.lean | 4 ++-- src/probability_theory/independence.lean | 3 +-- src/topology/metric_space/emetric_space.lean | 8 ++++---- 4 files changed, 8 insertions(+), 9 deletions(-) diff --git a/src/data/polynomial/inductions.lean b/src/data/polynomial/inductions.lean index c01a3ce77a882..aea91c1917668 100644 --- a/src/data/polynomial/inductions.lean +++ b/src/data/polynomial/inductions.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Scott Morrison, Damiano Testa, Jens Wagemaker -/ -import data.finset.intervals +import data.nat.intervals import data.polynomial.degree.definitions /-! diff --git a/src/data/polynomial/iterated_deriv.lean b/src/data/polynomial/iterated_deriv.lean index cff6f402f0522..dae05263e7fb7 100644 --- a/src/data/polynomial/iterated_deriv.lean +++ b/src/data/polynomial/iterated_deriv.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ -import data.finset.intervals +import data.nat.intervals import data.polynomial.derivative import tactic.linarith @@ -128,7 +128,7 @@ begin { intro m, rw [iterated_deriv_succ, coeff_derivative, ih (m+1), mul_right_comm], apply congr_arg2, { have set_eq : (Ico m.succ (m + k.succ.succ)) = (Ico (m + 1).succ (m + 1 + k.succ)) ∪ {m+1}, - { rw [union_comm, ←insert_eq, Ico.insert_succ_bot, add_succ, add_succ, add_succ _ k, + { rw [union_comm, ←insert_eq, Ico_insert_succ_left, add_succ, add_succ, add_succ _ k, ←succ_eq_add_one, succ_add], rw succ_eq_add_one, linarith }, diff --git a/src/probability_theory/independence.lean b/src/probability_theory/independence.lean index 1c265637afd3e..751db8ca4bff2 100644 --- a/src/probability_theory/independence.lean +++ b/src/probability_theory/independence.lean @@ -3,10 +3,9 @@ Copyright (c) 2021 Rémy Degenne. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ +import algebra.big_operators.intervals import measure_theory.measure.measure_space import measure_theory.pi_system -import algebra.big_operators.intervals -import data.finset.intervals /-! # Independence of sets of sets and measure spaces (σ-algebras) diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index b66e71cc72d9d..a094967479d4a 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -3,11 +3,11 @@ Copyright (c) 2015, 2017 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel -/ +import data.nat.intervals import data.real.ennreal -import data.finset.intervals -import topology.uniform_space.uniform_embedding import topology.uniform_space.pi import topology.uniform_space.uniform_convergence +import topology.uniform_space.uniform_embedding /-! # Extended metric spaces @@ -135,9 +135,9 @@ begin exact le_refl (0:ℝ≥0∞) }, { assume n hn hrec, calc edist (f m) (f (n+1)) ≤ edist (f m) (f n) + edist (f n) (f (n+1)) : edist_triangle _ _ _ - ... ≤ ∑ i in finset.Ico m n, _ + _ : add_le_add hrec (le_refl _) + ... ≤ ∑ i in finset.Ico m n, _ + _ : add_le_add hrec le_rfl ... = ∑ i in finset.Ico m (n+1), _ : - by rw [finset.Ico_insert_right hn, finset.sum_insert, add_comm]; simp } + by rw [nat.Ico_succ_right_eq_insert_Ico hn, finset.sum_insert, add_comm]; simp } end /-- The triangle (polygon) inequality for sequences of points; `finset.range` version. -/ From 3257c5ab1b44f6f229eaba35ec56081f0799c451 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sat, 18 Sep 2021 12:19:08 +0200 Subject: [PATCH 19/41] fix metric_space --- src/topology/instances/real.lean | 2 +- src/topology/metric_space/basic.lean | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/topology/instances/real.lean b/src/topology/instances/real.lean index 08e93c9936f88..7710b02d26f0d 100644 --- a/src/topology/instances/real.lean +++ b/src/topology/instances/real.lean @@ -66,7 +66,7 @@ instance : proper_space ℤ := ⟨ begin intros x r, rw closed_ball_eq, - exact (set.Icc_ℤ_finite _ _).is_compact, + exact (set.finite_Icc _ _).is_compact, end ⟩ end int diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 07b451fe0989b..a468d88cde752 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -196,7 +196,7 @@ begin calc dist (f m) (f (n+1)) ≤ dist (f m) (f n) + dist _ _ : dist_triangle _ _ _ ... ≤ ∑ i in finset.Ico m n, _ + _ : add_le_add hrec (le_refl _) ... = ∑ i in finset.Ico m (n+1), _ : - by rw [finset.Ico_insert_right hn, finset.sum_insert, add_comm]; simp } + by rw [nat.Ico_succ_right_eq_insert_Ico hn, finset.sum_insert, add_comm]; simp } end /-- The triangle (polygon) inequality for sequences of points; `finset.range` version. -/ @@ -1884,7 +1884,7 @@ begin refine tendsto_cocompact_of_tendsto_dist_comp_at_top (0 : ℝ) _, simp only [filter.tendsto_at_top, eventually_cofinite, not_le, ← mem_ball], change ∀ r : ℝ, finite (coe ⁻¹' (ball (0 : ℝ) r)), - simp [real.ball_eq, Ioo_ℤ_finite] + simp [real.ball_eq, set.finite_Ioo] end end int From 34755403c3890407d886417fb39684afbab5e9d1 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sat, 18 Sep 2021 12:22:57 +0200 Subject: [PATCH 20/41] really fix it --- src/topology/metric_space/basic.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index a468d88cde752..755f608d07b94 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel -/ -import topology.metric_space.emetric_space +import data.int.intervals import topology.algebra.ordered.basic -import data.nat.intervals +import topology.metric_space.emetric_space /-! # Metric spaces @@ -1886,7 +1886,7 @@ begin change ∀ r : ℝ, finite (coe ⁻¹' (ball (0 : ℝ) r)), simp [real.ball_eq, set.finite_Ioo] end - +#exit end int /-- We now define `metric_space`, extending `pseudo_metric_space`. -/ From 8ba820438cc4e5e2d0c69092252bc5bb261e6fb0 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sat, 18 Sep 2021 12:47:33 +0200 Subject: [PATCH 21/41] really really fix it --- src/topology/metric_space/basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 755f608d07b94..03b6281741e11 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -1884,9 +1884,9 @@ begin refine tendsto_cocompact_of_tendsto_dist_comp_at_top (0 : ℝ) _, simp only [filter.tendsto_at_top, eventually_cofinite, not_le, ← mem_ball], change ∀ r : ℝ, finite (coe ⁻¹' (ball (0 : ℝ) r)), - simp [real.ball_eq, set.finite_Ioo] + simp [real.ball_eq, set.finite_Ioo], end -#exit + end int /-- We now define `metric_space`, extending `pseudo_metric_space`. -/ From 1f27ad1093da79f1e454612a6569ca982fa50145 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sat, 18 Sep 2021 13:45:30 +0200 Subject: [PATCH 22/41] fix data.nat.totient --- src/data/nat/totient.lean | 12 ++---------- 1 file changed, 2 insertions(+), 10 deletions(-) diff --git a/src/data/nat/totient.lean b/src/data/nat/totient.lean index 3027ee6308b18..df081726edc43 100644 --- a/src/data/nat/totient.lean +++ b/src/data/nat/totient.lean @@ -181,16 +181,8 @@ begin { rintro rfl, rw [totient_one, nat.sub_self] at h, exact one_ne_zero h } }, - have hsplit : finset.range p = {0} ∪ (finset.Ico 1 p), - { rw [finset.range_eq_Ico, ←finset.Ico_union_Ico_eq_Ico zero_le_one hp.le, - finset.Ico_succ_singleton] }, - have hempty : finset.filter p.coprime {0} = ∅, - { simp only [finset.filter_singleton, nat.coprime_zero_right, hp.ne', if_false] }, - rw [totient_eq_card_coprime, hsplit, finset.filter_union, hempty, finset.empty_union, - ←nat.card_Ico 1 p] at h, - refine p.prime_of_coprime hp (λ n hn hnz, _), - apply finset.filter_card_eq h n, - refine finset.mem_Ico.mpr ⟨_, hn⟩, + rw [totient_eq_card_coprime, range_eq_Ico, ←Ico_insert_succ_left hp.le, finset.filter_insert, if_neg (tactic.norm_num.nat_coprime_helper_zero_right p hp), ←nat.card_Ico 1 p] at h, + refine p.prime_of_coprime hp (λ n hn hnz, finset.filter_card_eq h n $ finset.mem_Ico.mpr ⟨_, hn⟩), rwa [succ_le_iff, pos_iff_ne_zero], end From 9a1e00154c7832106d66ad1e7945de6113e7ddfe Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sat, 18 Sep 2021 14:49:56 +0200 Subject: [PATCH 23/41] add imports to interval_cases --- src/tactic/interval_cases.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/src/tactic/interval_cases.lean b/src/tactic/interval_cases.lean index 5da1d9006bc7c..f6e64285f0f66 100644 --- a/src/tactic/interval_cases.lean +++ b/src/tactic/interval_cases.lean @@ -26,7 +26,9 @@ The hypotheses should be in the form `hl : a ≤ n` and `hu : n < b`, in which case `interval_cases` calls `fin_cases` on the resulting fact `n ∈ set.Ico a b`. -/ -import data.nat.intervals +import data.fin.intervals +import data.int.intervals +import data.pnat.intervals import tactic.fin_cases open set From f834ac66e0045688e0397a53b6ffbbf459d5e843 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sat, 18 Sep 2021 17:42:05 +0200 Subject: [PATCH 24/41] fix topology.algebra.infinite_sum --- src/topology/algebra/infinite_sum.lean | 12 ++++-------- 1 file changed, 4 insertions(+), 8 deletions(-) diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index f1baeeea27b92..e50a2bad14aa4 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -4,14 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import algebra.big_operators.intervals -import topology.instances.real -import topology.algebra.module -import algebra.indicator_function -import data.equiv.encodable.lattice -import data.fintype.card -import data.nat.parity import algebra.big_operators.nat_antidiagonal -import order.filter.at_top_bot +import data.equiv.encodable.lattice +import topology.algebra.mul_action +import topology.instances.real /-! # Infinite sum over a topological monoid @@ -1163,7 +1159,7 @@ alias summable_abs_iff ↔ summable.of_abs summable.abs end linear_order section cauchy_seq -open finset.Ico filter +open filter /-- If the extended distance between consecutive points of a sequence is estimated by a summable series of `nnreal`s, then the original sequence is a Cauchy sequence. -/ From c01878a17054f93fe22c0d2e5dc3e0db55232c4e Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sat, 18 Sep 2021 20:22:38 +0200 Subject: [PATCH 25/41] fix topology.instances.nnreal --- src/topology/instances/nnreal.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/topology/instances/nnreal.lean b/src/topology/instances/nnreal.lean index ecc2a360c181e..b8b217a182038 100644 --- a/src/topology/instances/nnreal.lean +++ b/src/topology/instances/nnreal.lean @@ -96,7 +96,7 @@ lemma tendsto_of_real {f : filter α} {m : α → ℝ} {x : ℝ} (h : tendsto m (continuous_of_real.tendsto _).comp h lemma nhds_zero : 𝓝 (0 : ℝ≥0) = ⨅a ≠ 0, 𝓟 (Iio a) := -nhds_bot_order.trans $ by simp [bot_lt_iff_ne_bot, Iio] +nhds_bot_order.trans $ by simp [bot_lt_iff_ne_bot, set.Iio] lemma nhds_zero_basis : (𝓝 (0 : ℝ≥0)).has_basis (λ a : ℝ≥0, 0 < a) (λ a, Iio a) := nhds_bot_basis From 56c2a6c4fc26e4b5c06332e43e48d0a4516c7759 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sat, 18 Sep 2021 23:02:44 +0200 Subject: [PATCH 26/41] fixed unsigned_hahn --- src/measure_theory/decomposition/unsigned_hahn.lean | 12 +++++------- 1 file changed, 5 insertions(+), 7 deletions(-) diff --git a/src/measure_theory/decomposition/unsigned_hahn.lean b/src/measure_theory/decomposition/unsigned_hahn.lean index 1845327959e89..0a9af28265605 100644 --- a/src/measure_theory/decomposition/unsigned_hahn.lean +++ b/src/measure_theory/decomposition/unsigned_hahn.lean @@ -50,7 +50,8 @@ begin have to_nnreal_ν : ∀s, ((ν s).to_nnreal : ℝ≥0∞) = ν s := (assume s, ennreal.coe_to_nnreal $ ne_top_of_lt $ hν _), - have d_empty : d ∅ = 0, { simp [d], rw [measure_empty, measure_empty], simp }, + have d_empty : d ∅ = 0, + { change _ - _ = _, rw [measure_empty, measure_empty, sub_self] }, have d_split : ∀s t, measurable_set s → measurable_set t → d s = d (s \ t) + d (s ∩ t), @@ -114,16 +115,13 @@ begin { assume a b c d hab hcd, dsimp only [f], rw [finset.inf_eq_infi, finset.inf_eq_infi], - refine bInter_subset_bInter_left _, - simp, - rintros j ⟨hbj, hjc⟩, - exact ⟨le_trans hab hbj, lt_of_lt_of_le hjc $ add_le_add_right hcd 1⟩ }, + exact bInter_subset_bInter_left (finset.Ico_subset_Ico hab $ nat.succ_le_succ hcd) }, have f_succ : ∀n m, n ≤ m → f n (m + 1) = f n m ∩ e (m + 1), { assume n m hnm, have : n ≤ m + 1 := le_of_lt (nat.succ_le_succ hnm), simp only [f], - rw [finset.Ico_insert_right this, finset.inf_insert, set.inter_comm], + rw [nat.Ico_succ_right_eq_insert_Ico this, finset.inf_insert, set.inter_comm], refl }, have le_d_f : ∀n m, m ≤ n → γ - 2 * ((1 / 2) ^ m) + (1 / 2) ^ n ≤ d (f m n), @@ -131,7 +129,7 @@ begin refine nat.le_induction _ _ n h, { have := he₂ m, simp only [f], - rw [finset.Ico_succ_singleton, finset.inf_singleton], + rw [nat.Ico_succ_singleton, finset.inf_singleton], exact aux this }, { assume n (hmn : m ≤ n) ih, have : γ + (γ - 2 * (1 / 2)^m + (1 / 2) ^ (n + 1)) ≤ γ + d (f m (n + 1)), From 7be94b3f61fbc3e3b37b925c577d7c74a3054eec Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sun, 19 Sep 2021 08:04:29 +0200 Subject: [PATCH 27/41] fix analysis.analytic.inverse --- src/analysis/analytic/inverse.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/analysis/analytic/inverse.lean b/src/analysis/analytic/inverse.lean index c89e88a29aded..ca347a5d6dcd6 100644 --- a/src/analysis/analytic/inverse.lean +++ b/src/analysis/analytic/inverse.lean @@ -417,7 +417,7 @@ let I := ∥(i.symm : F →L[𝕜] E)∥ in calc ∑ k in Ico 1 (n + 1), a ^ k * ∥p.right_inv i k∥ = a * I + ∑ k in Ico 2 (n + 1), a ^ k * ∥p.right_inv i k∥ : by simp only [linear_isometry_equiv.norm_map, pow_one, right_inv_coeff_one, - Ico_succ_singleton, sum_singleton, ← sum_Ico_consecutive _ one_le_two hn] + nat.Ico_succ_singleton, sum_singleton, ← sum_Ico_consecutive _ one_le_two hn] ... = a * I + ∑ k in Ico 2 (n + 1), a ^ k * ∥(i.symm : F →L[𝕜] E).comp_continuous_multilinear_map (∑ c in ({c | 1 < composition.length c}.to_finset : finset (composition k)), From b68044617fecd06fbaf9f0f4dca0d5ddf5608a1d Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sun, 19 Sep 2021 08:18:04 +0200 Subject: [PATCH 28/41] add namespaces to aliases --- src/data/finset/intervals.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/data/finset/intervals.lean b/src/data/finset/intervals.lean index eb4d4db969454..86133763f78f0 100644 --- a/src/data/finset/intervals.lean +++ b/src/data/finset/intervals.lean @@ -51,9 +51,9 @@ by rw [←coe_eq_empty, coe_Ioc, set.Ioc_eq_empty_iff] @[simp] lemma Ioo_eq_empty_iff [densely_ordered α] : Ioo a b = ∅ ↔ ¬a < b := by rw [←coe_eq_empty, coe_Ioo, set.Ioo_eq_empty_iff] -alias Icc_eq_empty_iff ↔ _ Icc_eq_empty -alias Ico_eq_empty_iff ↔ _ Ico_eq_empty -alias Ioc_eq_empty_iff ↔ _ Ioc_eq_empty +alias Icc_eq_empty_iff ↔ _ finset.Icc_eq_empty +alias Ico_eq_empty_iff ↔ _ finset.Ico_eq_empty +alias Ioc_eq_empty_iff ↔ _ finset.Ioc_eq_empty @[simp] lemma Ioo_eq_empty (h : ¬a < b) : Ioo a b = ∅ := eq_empty_iff_forall_not_mem.2 $ λ x hx, h ((mem_Ioo.1 hx).1.trans (mem_Ioo.1 hx).2) From 78bf5d5511b240eedfd78a117c3aeb2360f0a38a Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sun, 19 Sep 2021 10:06:06 +0200 Subject: [PATCH 29/41] fix exp_log --- src/analysis/special_functions/exp_log.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/analysis/special_functions/exp_log.lean b/src/analysis/special_functions/exp_log.lean index 4c0575ecc9d80..6251becec22c3 100644 --- a/src/analysis/special_functions/exp_log.lean +++ b/src/analysis/special_functions/exp_log.lean @@ -706,9 +706,9 @@ begin eventually_at_top.1 ((tendsto_pow_const_div_const_pow_of_one_lt n (one_lt_exp_iff.2 zero_lt_one)).eventually (gt_mem_nhds this)), simp only [← exp_nat_mul, mul_one, div_lt_iff, exp_pos, ← div_eq_inv_mul] at hN, - refine ⟨N, trivial, λ x hx, _⟩, rw mem_Ioi at hx, + refine ⟨N, trivial, λ x hx, _⟩, rw set.mem_Ioi at hx, have hx₀ : 0 < x, from N.cast_nonneg.trans_lt hx, - rw [mem_Ici, le_div_iff (pow_pos hx₀ _), ← le_div_iff' hC₀], + rw [set.mem_Ici, le_div_iff (pow_pos hx₀ _), ← le_div_iff' hC₀], calc x ^ n ≤ ⌈x⌉₊ ^ n : pow_le_pow_of_le_left hx₀.le (le_nat_ceil _) _ ... ≤ exp ⌈x⌉₊ / (exp 1 * C) : (hN _ (lt_nat_ceil.2 hx).le).le ... ≤ exp (x + 1) / (exp 1 * C) : div_le_div_of_le (mul_pos (exp_pos _) hC₀).le From 18ba02f87d2e437033e2ccc6fe5515f0eb300b91 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Sun, 19 Sep 2021 14:58:22 +0200 Subject: [PATCH 30/41] fix test --- src/data/nat/totient.lean | 5 +++-- src/logic/basic.lean | 3 --- test/fin_cases.lean | 4 ++-- 3 files changed, 5 insertions(+), 7 deletions(-) diff --git a/src/data/nat/totient.lean b/src/data/nat/totient.lean index df081726edc43..08321d2317cc2 100644 --- a/src/data/nat/totient.lean +++ b/src/data/nat/totient.lean @@ -10,7 +10,7 @@ import data.zmod.basic /-! # Euler's totient function -This file defines [Euler's totient function][https://en.wikipedia.org/wiki/Euler's_totient_function] +This file defines [Euler's totient function](https://en.wikipedia.org/wiki/Euler's_totient_function) `nat.totient n` which counts the number of naturals less than `n` that are coprime with `n`. We prove the divisor sum formula, namely that `n` equals `φ` summed over the divisors of `n`. See `sum_totient`. We also prove two lemmas to help compute totients, namely `totient_mul` and @@ -181,7 +181,8 @@ begin { rintro rfl, rw [totient_one, nat.sub_self] at h, exact one_ne_zero h } }, - rw [totient_eq_card_coprime, range_eq_Ico, ←Ico_insert_succ_left hp.le, finset.filter_insert, if_neg (tactic.norm_num.nat_coprime_helper_zero_right p hp), ←nat.card_Ico 1 p] at h, + rw [totient_eq_card_coprime, range_eq_Ico, ←Ico_insert_succ_left hp.le, finset.filter_insert, + if_neg (tactic.norm_num.nat_coprime_helper_zero_right p hp), ←nat.card_Ico 1 p] at h, refine p.prime_of_coprime hp (λ n hn hnz, finset.filter_card_eq h n $ finset.mem_Ico.mpr ⟨_, hn⟩), rwa [succ_le_iff, pos_iff_ne_zero], end diff --git a/src/logic/basic.lean b/src/logic/basic.lean index cf6022a90105c..faa9c7c9dadf4 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -161,9 +161,6 @@ lemma ne_comm {α} {a b : α} : a ≠ b ↔ b ≠ a := ⟨ne.symm, ne.symm⟩ (∀ {c}, a = c ↔ b = c) ↔ (a = b) := ⟨λ h, by rw h, λ h a, by rw h⟩ -/-- Annotation identity function to force Lean to see `a` in the type of the goal. -/ -abbreviation id_annotate {α} (a : α) {β} (b : β) := b - /-- Wrapper for adding elementary propositions to the type class systems. Warning: this can easily be abused. See the rest of this docstring for details. diff --git a/test/fin_cases.lean b/test/fin_cases.lean index 818d1ee2ac360..cf381372a47ed 100644 --- a/test/fin_cases.lean +++ b/test/fin_cases.lean @@ -3,10 +3,10 @@ Copyright (c) 2018 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import tactic.fin_cases +import data.nat.intervals import data.nat.prime import group_theory.perm.sign -import data.finset.intervals +import tactic.fin_cases example (f : ℕ → Prop) (p : fin 3) (h0 : f 0) (h1 : f 1) (h2 : f 2) : f p.val := begin From 10eae5ce43888172ef67b66e00c84e180aa8a5fb Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Tue, 21 Sep 2021 16:55:12 +0200 Subject: [PATCH 31/41] delete mwe and add subsingleton instance --- src/data/finset/mwe.lean | 0 src/order/locally_finite.lean | 9 +++++++++ 2 files changed, 9 insertions(+) delete mode 100644 src/data/finset/mwe.lean diff --git a/src/data/finset/mwe.lean b/src/data/finset/mwe.lean deleted file mode 100644 index e69de29bb2d1d..0000000000000 diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 0be8d01f55cca..05c11608d2a3a 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -420,6 +420,15 @@ noncomputable def fintype.to_locally_finite_order [fintype α] : finset_mem_Ioc := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioc], finset_mem_Ioo := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioo] } +instance : subsingleton (locally_finite_order α) := +subsingleton.intro _ (λ h₀ h₁, begin + ext, + { rw [h₀.finset_mem_Icc, h₁.finset_mem_Icc] }, + { rw [h₀.finset_mem_Ico, h₁.finset_mem_Ico] }, + { rw [h₀.finset_mem_Ioc, h₁.finset_mem_Ioc] }, + { rw [h₀.finset_mem_Ioo, h₁.finset_mem_Ioo] } +end) + variables [preorder β] [locally_finite_order β] -- Should this be called `locally_finite_order.lift`? From 6c36a479408ca17d7dd34ff47823a72124c22329 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Tue, 21 Sep 2021 20:21:24 +0200 Subject: [PATCH 32/41] fix subsingleton instance --- src/order/locally_finite.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 05c11608d2a3a..35a24449baa64 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -95,7 +95,7 @@ open finset /-- A locally finite order is an order where bounded intervals are finite. When you don't care too much about definitional equality, you can use `locally_finite_order.of_Icc` or `locally_finite_order.of_finite_Icc` to build a locally finite order from just `finset.Icc`. -/ -class locally_finite_order (α : Type*) [preorder α] := +@[ext] class locally_finite_order (α : Type*) [preorder α] := (finset_Icc : α → α → finset α) (finset_Ico : α → α → finset α) (finset_Ioc : α → α → finset α) @@ -421,7 +421,7 @@ noncomputable def fintype.to_locally_finite_order [fintype α] : finset_mem_Ioo := λ a b x, by rw [set.finite.mem_to_finset, set.mem_Ioo] } instance : subsingleton (locally_finite_order α) := -subsingleton.intro _ (λ h₀ h₁, begin +subsingleton.intro (λ h₀ h₁, begin ext, { rw [h₀.finset_mem_Icc, h₁.finset_mem_Icc] }, { rw [h₀.finset_mem_Ico, h₁.finset_mem_Ico] }, From 990ba78528d86b97cf6b7a9de1ba4faeb7d282f8 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Wed, 22 Sep 2021 08:54:05 +0200 Subject: [PATCH 33/41] get rid of the @[ext] --- src/order/locally_finite.lean | 18 ++++++++++++------ 1 file changed, 12 insertions(+), 6 deletions(-) diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 35a24449baa64..92449d794d3cf 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -95,7 +95,7 @@ open finset /-- A locally finite order is an order where bounded intervals are finite. When you don't care too much about definitional equality, you can use `locally_finite_order.of_Icc` or `locally_finite_order.of_finite_Icc` to build a locally finite order from just `finset.Icc`. -/ -@[ext] class locally_finite_order (α : Type*) [preorder α] := +class locally_finite_order (α : Type*) [preorder α] := (finset_Icc : α → α → finset α) (finset_Ico : α → α → finset α) (finset_Ioc : α → α → finset α) @@ -422,11 +422,17 @@ noncomputable def fintype.to_locally_finite_order [fintype α] : instance : subsingleton (locally_finite_order α) := subsingleton.intro (λ h₀ h₁, begin - ext, - { rw [h₀.finset_mem_Icc, h₁.finset_mem_Icc] }, - { rw [h₀.finset_mem_Ico, h₁.finset_mem_Ico] }, - { rw [h₀.finset_mem_Ioc, h₁.finset_mem_Ioc] }, - { rw [h₀.finset_mem_Ioo, h₁.finset_mem_Ioo] } + cases h₀, + cases h₁, + have hIcc : h₀_finset_Icc = h₁_finset_Icc, + { ext a b x, rw [h₀_finset_mem_Icc, h₁_finset_mem_Icc] }, + have hIco : h₀_finset_Ico = h₁_finset_Ico, + { ext a b x, rw [h₀_finset_mem_Ico, h₁_finset_mem_Ico] }, + have hIoc : h₀_finset_Ioc = h₁_finset_Ioc, + { ext a b x, rw [h₀_finset_mem_Ioc, h₁_finset_mem_Ioc] }, + have hIoo : h₀_finset_Ioo = h₁_finset_Ioo, + { ext a b x, rw [h₀_finset_mem_Ioo, h₁_finset_mem_Ioo] }, + simp_rw [hIcc, hIco, hIoc, hIoo], end) variables [preorder β] [locally_finite_order β] From f0d3a8a53628855ffe0cb5a2436774113a633163 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Wed, 22 Sep 2021 14:06:42 +0200 Subject: [PATCH 34/41] fix lint --- src/data/list/erase_dup.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/list/erase_dup.lean b/src/data/list/erase_dup.lean index 6cf545f624532..a4656bfbb9b8a 100644 --- a/src/data/list/erase_dup.lean +++ b/src/data/list/erase_dup.lean @@ -56,7 +56,7 @@ theorem nodup_erase_dup : ∀ l : list α, nodup (erase_dup l) := pairwise_pw_fi theorem erase_dup_eq_self {l : list α} : erase_dup l = l ↔ nodup l := pw_filter_eq_self -@[simp] protected lemma nodup.erase_dup {l : list α} (h : l.nodup) : l.erase_dup = l := +protected lemma nodup.erase_dup {l : list α} (h : l.nodup) : l.erase_dup = l := list.erase_dup_eq_self.2 h @[simp] theorem erase_dup_idempotent {l : list α} : erase_dup (erase_dup l) = erase_dup l := From 8432cfc04bf0acfa09a8574218b6e8baca3231e4 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Mon, 4 Oct 2021 22:12:55 +0200 Subject: [PATCH 35/41] bump mathlib --- src/algebra/char_zero.lean | 4 + src/data/fin/interval.lean | 15 +- src/data/fin/intervals.lean | 65 ------- .../int/{intervals.lean => interval.lean} | 72 +++++--- src/data/nat/interval.lean | 93 +++++++++- src/data/nat/intervals.lean | 171 ------------------ src/data/pnat/interval.lean | 10 + src/data/pnat/intervals.lean | 63 ------- 8 files changed, 164 insertions(+), 329 deletions(-) delete mode 100644 src/data/fin/intervals.lean rename src/data/int/{intervals.lean => interval.lean} (55%) delete mode 100644 src/data/nat/intervals.lean delete mode 100644 src/data/pnat/intervals.lean diff --git a/src/algebra/char_zero.lean b/src/algebra/char_zero.lean index a775dc0f6aa41..03238d7319d72 100644 --- a/src/algebra/char_zero.lean +++ b/src/algebra/char_zero.lean @@ -65,6 +65,10 @@ variables {R : Type*} [add_monoid R] [has_one R] [char_zero R] theorem cast_injective : function.injective (coe : ℕ → R) := char_zero.cast_injective +/-- `nat.cast` as an embedding into monoids of characteristic `0`. -/ +@[simps] +def cast_embedding : ℕ ↪ R := ⟨coe, cast_injective⟩ + @[simp, norm_cast] theorem cast_inj {m n : ℕ} : (m : R) = n ↔ m = n := cast_injective.eq_iff diff --git a/src/data/fin/interval.lean b/src/data/fin/interval.lean index 95d8fb0c2b96a..84bfd9df0524e 100644 --- a/src/data/fin/interval.lean +++ b/src/data/fin/interval.lean @@ -10,10 +10,6 @@ import data.nat.interval This file proves that `fin n` is a `locally_finite_order` and calculates the cardinality of its intervals as finsets and fintypes. - -## TODO - -@Yaël: Add `finset.Ico`. Coming very soon. -/ open finset fin @@ -26,6 +22,7 @@ namespace fin variables {n} (a b : fin n) lemma Icc_eq_finset_subtype : Icc a b = (Icc (a : ℕ) b).subtype (λ x, x < n) := rfl +lemma Ico_eq_finset_subtype : Ico a b = (Ico (a : ℕ) b).subtype (λ x, x < n) := rfl lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).subtype (λ x, x < n) := rfl lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).subtype (λ x, x < n) := rfl @@ -33,6 +30,10 @@ lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).subtype (λ x, x < n) (Icc a b).map (function.embedding.subtype _) = Icc (a : ℕ) b := map_subtype_embedding_Icc _ _ _ (λ _ c x _ hx _, hx.trans_lt) +@[simp] lemma map_subtype_embedding_Ico : + (Ico a b).map (function.embedding.subtype _) = Ico (a : ℕ) b := +map_subtype_embedding_Ico _ _ _ (λ _ c x _ hx _, hx.trans_lt) + @[simp] lemma map_subtype_embedding_Ioc : (Ioc a b).map (function.embedding.subtype _) = Ioc (a : ℕ) b := map_subtype_embedding_Ioc _ _ _ (λ _ c x _ hx _, hx.trans_lt) @@ -44,6 +45,9 @@ map_subtype_embedding_Ioo _ _ _ (λ _ c x _ hx _, hx.trans_lt) @[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := by rw [←nat.card_Icc, ←map_subtype_embedding_Icc, card_map] +@[simp] lemma card_Ico : (Ico a b).card = b - a := +by rw [←nat.card_Ico, ←map_subtype_embedding_Ico, card_map] + @[simp] lemma card_Ioc : (Ioc a b).card = b - a := by rw [←nat.card_Ioc, ←map_subtype_embedding_Ioc, card_map] @@ -53,6 +57,9 @@ by rw [←nat.card_Ioo, ←map_subtype_embedding_Ioo, card_map] @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := by rw [←card_Icc, fintype.card_of_finset] +@[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = b - a := +by rw [←card_Ico, fintype.card_of_finset] + @[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = b - a := by rw [←card_Ioc, fintype.card_of_finset] diff --git a/src/data/fin/intervals.lean b/src/data/fin/intervals.lean deleted file mode 100644 index 650ec889d18da..0000000000000 --- a/src/data/fin/intervals.lean +++ /dev/null @@ -1,65 +0,0 @@ -/- -Copyright (c) 2020 Scott Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison, Yaël Dillies --/ -import data.nat.intervals - -/-! -# Finite intervals in `fin n` - -This file proves that `fin n` is a `locally_finite_order` and calculates the cardinality of its -intervals as finsets and fintypes. --/ - -open finset fin - -variables (n : ℕ) - -instance : locally_finite_order (fin n) := subtype.locally_finite_order _ - -namespace fin -variables {n} (a b : fin n) - -lemma Icc_eq_finset_subtype : Icc a b = (Icc (a : ℕ) b).subtype (λ x, x < n) := rfl -lemma Ico_eq_finset_subtype : Ico a b = (Ico (a : ℕ) b).subtype (λ x, x < n) := rfl -lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).subtype (λ x, x < n) := rfl -lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).subtype (λ x, x < n) := rfl - -lemma map_subtype_embedding_Icc : (Icc a b).map (function.embedding.subtype _) = Icc (a : ℕ) b := -map_subtype_embedding_Icc _ _ _ (λ _ c x _ hx _, hx.trans_lt) - -lemma map_subtype_embedding_Ico : (Ico a b).map (function.embedding.subtype _) = Ico (a : ℕ) b := -map_subtype_embedding_Ico _ _ _ (λ _ c x _ hx _, hx.trans_lt) - -lemma map_subtype_embedding_Ioc : (Ioc a b).map (function.embedding.subtype _) = Ioc (a : ℕ) b := -map_subtype_embedding_Ioc _ _ _ (λ _ c x _ hx _, hx.trans_lt) - -lemma map_subtype_embedding_Ioo : (Ioo a b).map (function.embedding.subtype _) = Ioo (a : ℕ) b := -map_subtype_embedding_Ioo _ _ _ (λ _ c x _ hx _, hx.trans_lt) - -@[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := -by rw [←nat.card_Icc, ←map_subtype_embedding_Icc, card_map] - -@[simp] lemma card_Ico : (Ico a b).card = b - a := -by rw [←nat.card_Ico, ←map_subtype_embedding_Ico, card_map] - -@[simp] lemma card_Ioc : (Ioc a b).card = b - a := -by rw [←nat.card_Ioc, ←map_subtype_embedding_Ioc, card_map] - -@[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := -by rw [←nat.card_Ioo, ←map_subtype_embedding_Ioo, card_map] - -@[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := -by rw [←card_Icc, fintype.card_of_finset] - -@[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = b - a := -by rw [←card_Ico, fintype.card_of_finset] - -@[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = b - a := -by rw [←card_Ioc, fintype.card_of_finset] - -@[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := -by rw [←card_Ioo, fintype.card_of_finset] - -end fin diff --git a/src/data/int/intervals.lean b/src/data/int/interval.lean similarity index 55% rename from src/data/int/intervals.lean rename to src/data/int/interval.lean index 13bb7a54862af..fc353f4abc483 100644 --- a/src/data/int/intervals.lean +++ b/src/data/int/interval.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import data.int.basic -import data.nat.intervals +import data.nat.interval /-! # Finite intervals of integers @@ -14,34 +14,32 @@ intervals as finsets and fintypes. -/ open finset int -private lemma inj (a : ℤ) : function.injective (λ n : ℕ, (n : ℤ) + a) := -λ m n, by { rw [add_left_inj, coe_nat_inj'], exact id } instance : locally_finite_order ℤ := -{ finset_Icc := λ a b, (finset.range (b + 1 - a).to_nat).map - ⟨λ n, n + a, inj a⟩, - finset_Ico := λ a b, (finset.range (b - a).to_nat).map - ⟨λ n, n + a, inj a⟩, - finset_Ioc := λ a b, (finset.range (b - a).to_nat).map - ⟨λ n, n + (a + 1), inj (a + 1)⟩, - finset_Ioo := λ a b, (finset.range (b - a - 1).to_nat).map - ⟨λ n, n + (a + 1), inj (a + 1)⟩, +{ finset_Icc := λ a b, (finset.range (b + 1 - a).to_nat).map $ + nat.cast_embedding.trans $ add_left_embedding a, + finset_Ico := λ a b, (finset.range (b - a).to_nat).map $ + nat.cast_embedding.trans $ add_left_embedding a, + finset_Ioc := λ a b, (finset.range (b - a).to_nat).map $ + nat.cast_embedding.trans $ add_left_embedding (a + 1), + finset_Ioo := λ a b, (finset.range (b - a - 1).to_nat).map $ + nat.cast_embedding.trans $ add_left_embedding (a + 1), finset_mem_Icc := λ a b x, begin - simp only [int.lt_to_nat, exists_prop, mem_range, add_comm, function.embedding.coe_fn_mk, - mem_map], + simp_rw [mem_map, exists_prop, mem_range, int.lt_to_nat, function.embedding.trans_apply, + nat.cast_embedding_apply, add_left_embedding_apply, nat_cast_eq_coe_nat], split, { rintro ⟨a, h, rfl⟩, rw [lt_sub_iff_add_lt, int.lt_add_one_iff, add_comm] at h, - refine ⟨int.le.intro rfl, h⟩ }, + exact ⟨int.le.intro rfl, h⟩ }, { rintro ⟨ha, hb⟩, use (x - a).to_nat, - rw to_nat_sub_of_le ha, rw ←lt_add_one_iff at hb, + rw to_nat_sub_of_le ha, exact ⟨sub_lt_sub_right hb _, add_sub_cancel'_right _ _⟩ } end, finset_mem_Ico := λ a b x, begin - simp only [int.lt_to_nat, exists_prop, mem_range, add_comm, function.embedding.coe_fn_mk, - mem_map], + simp_rw [mem_map, exists_prop, mem_range, int.lt_to_nat, function.embedding.trans_apply, + nat.cast_embedding_apply, add_left_embedding_apply, nat_cast_eq_coe_nat], split, { rintro ⟨a, h, rfl⟩, exact ⟨int.le.intro rfl, lt_sub_iff_add_lt'.mp h⟩ }, @@ -51,24 +49,24 @@ instance : locally_finite_order ℤ := exact ⟨sub_lt_sub_right hb _, add_sub_cancel'_right _ _⟩ } end, finset_mem_Ioc := λ a b x, begin - simp only [int.lt_to_nat, exists_prop, mem_range, add_comm, function.embedding.coe_fn_mk, - mem_map], + simp_rw [mem_map, exists_prop, mem_range, int.lt_to_nat, function.embedding.trans_apply, + nat.cast_embedding_apply, add_left_embedding_apply, nat_cast_eq_coe_nat], split, { rintro ⟨a, h, rfl⟩, - refine ⟨int.le.intro rfl, _⟩, - rwa [←add_one_le_iff, le_sub_iff_add_le', add_comm _ (1 : ℤ), ←add_assoc] at h }, + rw [←add_one_le_iff, le_sub_iff_add_le', add_comm _ (1 : ℤ), ←add_assoc] at h, + exact ⟨int.le.intro rfl, h⟩ }, { rintro ⟨ha, hb⟩, use (x - (a + 1)).to_nat, rw [to_nat_sub_of_le ha, ←add_one_le_iff, sub_add, add_sub_cancel], exact ⟨sub_le_sub_right hb _, add_sub_cancel'_right _ _⟩ } end, finset_mem_Ioo := λ a b x, begin - simp only [int.lt_to_nat, exists_prop, mem_range, add_comm, function.embedding.coe_fn_mk, - mem_map], + simp_rw [mem_map, exists_prop, mem_range, int.lt_to_nat, function.embedding.trans_apply, + nat.cast_embedding_apply, add_left_embedding_apply, nat_cast_eq_coe_nat], split, { rintro ⟨a, h, rfl⟩, - refine ⟨int.le.intro rfl, _⟩, - rwa [sub_sub, lt_sub_iff_add_lt'] at h }, + rw [sub_sub, lt_sub_iff_add_lt'] at h, + exact ⟨int.le.intro rfl, h⟩ }, { rintro ⟨ha, hb⟩, use (x - (a + 1)).to_nat, rw [to_nat_sub_of_le ha, sub_sub], @@ -98,6 +96,18 @@ by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } @[simp] lemma card_Ioo : (Ioo a b).card = (b - a - 1).to_nat := by { change (finset.map _ _).card = _, rw [finset.card_map, finset.card_range] } +lemma card_Icc_of_le (h : a ≤ b + 1) : ((Icc a b).card : ℤ) = b + 1 - a := +by rw [card_Icc, to_nat_sub_of_le h] + +lemma card_Ico_of_le (h : a ≤ b) : ((Ico a b).card : ℤ) = b - a := +by rw [card_Ico, to_nat_sub_of_le h] + +lemma card_Ioc_of_le (h : a ≤ b) : ((Ioc a b).card : ℤ) = b - a := +by rw [card_Ioc, to_nat_sub_of_le h] + +lemma card_Ioo_of_lt (h : a < b) : ((Ioo a b).card : ℤ) = b - a - 1 := +by rw [card_Ioo, sub_sub, to_nat_sub_of_le h] + @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = (b + 1 - a).to_nat := by rw [←card_Icc, fintype.card_of_finset] @@ -110,4 +120,16 @@ by rw [←card_Ioc, fintype.card_of_finset] @[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = (b - a - 1).to_nat := by rw [←card_Ioo, fintype.card_of_finset] +lemma card_fintype_Icc_of_le (h : a ≤ b + 1) : (fintype.card (set.Icc a b) : ℤ) = b + 1 - a := +by rw [card_fintype_Icc, to_nat_sub_of_le h] + +lemma card_fintype_Ico_of_le (h : a ≤ b) : (fintype.card (set.Ico a b) : ℤ) = b - a := +by rw [card_fintype_Ico, to_nat_sub_of_le h] + +lemma card_fintype_Ioc_of_le (h : a ≤ b) : (fintype.card (set.Ioc a b) : ℤ) = b - a := +by rw [card_fintype_Ioc, to_nat_sub_of_le h] + +lemma card_fintype_Ioo_of_lt (h : a < b) : (fintype.card (set.Ioo a b) : ℤ) = b - a - 1 := +by rw [card_fintype_Ioo, sub_sub, to_nat_sub_of_le h] + end int diff --git a/src/data/nat/interval.lean b/src/data/nat/interval.lean index d6ffe65ca9d05..492a8b82f14bd 100644 --- a/src/data/nat/interval.lean +++ b/src/data/nat/interval.lean @@ -13,7 +13,8 @@ intervals as finsets and fintypes. ## TODO -Add all the `finset.Ico` stuff. +Some lemmas can be generalized using `ordered_group`, `canonically_ordered_monoid` or `succ_order` +and subsequently be moved upstream to `data.finset.interval`. -/ open finset nat @@ -56,12 +57,24 @@ namespace nat variables (a b c : ℕ) lemma Icc_eq_range' : Icc a b = (list.range' a (b + 1 - a)).to_finset := rfl +lemma Ico_eq_range' : Ico a b = (list.range' a (b - a)).to_finset := rfl lemma Ioc_eq_range' : Ioc a b = (list.range' (a + 1) (b - a)).to_finset := rfl lemma Ioo_eq_range' : Ioo a b = (list.range' (a + 1) (b - a - 1)).to_finset := rfl +lemma Iio_eq_range : Iio = range := +by { ext b x, rw [mem_Iio, mem_range] } + +lemma Ico_zero_eq_range : Ico 0 a = range a := +by rw [←bot_eq_zero, ←Iio_eq_Ico, Iio_eq_range] + +lemma _root_.finset.range_eq_Ico : range = Ico 0 := by { funext, exact (Ico_zero_eq_range n).symm } + @[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := by rw [Icc_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] +@[simp] lemma card_Ico : (Ico a b).card = b - a := +by rw [Ico_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] + @[simp] lemma card_Ioc : (Ioc a b).card = b - a := by rw [Ioc_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] @@ -71,6 +84,9 @@ by rw [Ioo_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, li @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := by rw [←card_Icc, fintype.card_of_finset] +@[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = b - a := +by rw [←card_Ico, fintype.card_of_finset] + @[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = b - a := by rw [←card_Ioc, fintype.card_of_finset] @@ -82,4 +98,79 @@ by rw [←card_Ioo, fintype.card_of_finset] lemma Icc_succ_left : Icc a.succ b = Ioc a b := by { ext x, rw [mem_Icc, mem_Ioc, succ_le_iff] } +lemma Ico_succ_right : Ico a b.succ = Icc a b := +by { ext x, rw [mem_Ico, mem_Icc, lt_succ_iff] } + +lemma Ico_succ_left : Ico a.succ b = Ioo a b := +by { ext x, rw [mem_Ico, mem_Ioo, succ_le_iff] } + +lemma Icc_pred_right {b : ℕ} (h : 0 < b) : Icc a (b - 1) = Ico a b := +by { ext x, rw [mem_Icc, mem_Ico, lt_iff_le_pred h] } + +lemma Ico_succ_succ : Ico a.succ b.succ = Ioc a b := +by { ext x, rw [mem_Ico, mem_Ioc, succ_le_iff, lt_succ_iff] } + +@[simp] lemma Ico_succ_singleton : Ico a (a + 1) = {a} := +by rw [Ico_succ_right, Icc_self] + +@[simp] lemma Ico_pred_singleton {a : ℕ} (h : 0 < a) : Ico (a - 1) a = {a - 1} := +by rw [←Icc_pred_right _ h, Icc_self] + +variables {a b c} + +lemma Ico_succ_right_eq_insert_Ico (h : a ≤ b) : Ico a (b + 1) = insert b (Ico a b) := +by rw [Ico_succ_right, ←Ico_insert_right h] + +lemma Ico_insert_succ_left (h : a < b) : insert a (Ico a.succ b) = Ico a b := +by rw [Ico_succ_left, ←Ioo_insert_left h] + +lemma image_sub_const_Ico (h : c ≤ a) : + (Ico a b).image (λ x, x - c) = Ico (a - c) (b - c) := +begin + ext x, + rw mem_image, + split, + { rintro ⟨x, hx, rfl⟩, + rw [mem_Ico] at ⊢ hx, + exact ⟨sub_le_sub_right' hx.1 _, sub_lt_sub_right_of_le (h.trans hx.1) hx.2⟩ }, + { rintro h, + refine ⟨x + c, _, add_sub_cancel_right⟩, + rw mem_Ico at ⊢ h, + exact ⟨sub_le_iff_right.1 h.1, nat.add_lt_of_lt_sub_right h.2⟩ } +end + +lemma Ico_image_const_sub_eq_Ico (hac : a ≤ c) : + (Ico a b).image (λ x, c - x) = Ico (c + 1 - b) (c + 1 - a) := +begin + ext x, + rw [mem_image, mem_Ico], + split, + { rintro ⟨x, hx, rfl⟩, + rw mem_Ico at hx, + refine ⟨_, ((nat.sub_le_sub_left_iff hac).2 hx.1).trans_lt ((nat.sub_lt_sub_right_iff hac).2 + (nat.lt_succ_self _))⟩, + cases lt_or_le c b, + { rw nat.sub_eq_zero_of_le h, + exact zero_le _ }, + { rw ←succ_sub_succ c, + exact (nat.sub_le_sub_left_iff (succ_le_succ $ hx.2.le.trans h)).2 hx.2 } }, + { rintro ⟨hb, ha⟩, + rw [nat.lt_sub_left_iff_add_lt, lt_succ_iff] at ha, + have hx : x ≤ c := (nat.le_add_left _ _).trans ha, + refine ⟨c - x, _, nat.sub_sub_self hx⟩, + { rw mem_Ico, + refine ⟨nat.le_sub_right_of_add_le ha, _⟩, + rw [nat.sub_lt_left_iff_lt_add hx, ←succ_le_iff], + exact nat.le_add_of_sub_le_right hb } } +end + +lemma range_image_pred_top_sub (n : ℕ) : + (finset.range n).image (λ j, n - 1 - j) = finset.range n := +begin + cases n, + { rw [range_zero, image_empty] }, + { rw [finset.range_eq_Ico, Ico_image_const_sub_eq_Ico (zero_le _)], + simp_rw [succ_sub_succ, nat.sub_zero, nat.sub_self] } +end + end nat diff --git a/src/data/nat/intervals.lean b/src/data/nat/intervals.lean deleted file mode 100644 index 2f13479164638..0000000000000 --- a/src/data/nat/intervals.lean +++ /dev/null @@ -1,171 +0,0 @@ -/- -Copyright (c) 2021 Yaël Dillies. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies --/ -import data.finset.intervals - -/-! -# Finite intervals of naturals - -This file proves that `ℕ` is a `locally_finite_order` and calculates the cardinality of its -intervals as finsets and fintypes. --/ - -open finset nat - -instance : locally_finite_order ℕ := -{ finset_Icc := λ a b, (list.range' a (b + 1 - a)).to_finset, - finset_Ico := λ a b, (list.range' a (b - a)).to_finset, - finset_Ioc := λ a b, (list.range' (a + 1) (b - a)).to_finset, - finset_Ioo := λ a b, (list.range' (a + 1) (b - a - 1)).to_finset, - finset_mem_Icc := λ a b x, begin - rw [list.mem_to_finset, list.mem_range'], - cases le_or_lt a b, - { rw [nat.add_sub_cancel' (nat.lt_succ_of_le h).le, nat.lt_succ_iff] }, - { rw [nat.sub_eq_zero_iff_le.2 h, add_zero], - exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)) } - end, - finset_mem_Ico := λ a b x, begin - rw [list.mem_to_finset, list.mem_range'], - cases le_or_lt a b, - { rw [nat.add_sub_cancel' h] }, - { rw [nat.sub_eq_zero_iff_le.2 h.le, add_zero], - exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2.le)) } - end, - finset_mem_Ioc := λ a b x, begin - rw [list.mem_to_finset, list.mem_range'], - cases le_or_lt a b, - { rw [←succ_sub_succ, nat.add_sub_cancel' (succ_le_succ h), nat.lt_succ_iff, nat.succ_le_iff] }, - { rw [nat.sub_eq_zero_iff_le.2 h.le, add_zero], - exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.le.trans hx.2)) } - end, - finset_mem_Ioo := λ a b x, begin - rw [list.mem_to_finset, list.mem_range', nat.sub_sub], - cases le_or_lt (a + 1) b, - { rw [nat.add_sub_cancel' h, nat.succ_le_iff] }, - { rw [nat.sub_eq_zero_iff_le.2 h.le, add_zero], - exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)) } - end } - -namespace nat -variables (a b c : ℕ) - -lemma Icc_eq_range' : Icc a b = (list.range' a (b + 1 - a)).to_finset := rfl -lemma Ico_eq_range' : Ico a b = (list.range' a (b - a)).to_finset := rfl -lemma Ioc_eq_range' : Ioc a b = (list.range' (a + 1) (b - a)).to_finset := rfl -lemma Ioo_eq_range' : Ioo a b = (list.range' (a + 1) (b - a - 1)).to_finset := rfl - -lemma Iio_eq_range : Iio = range := -by { ext b x, rw [mem_Iio, mem_range] } - -lemma Ico_zero_eq_range : Ico 0 a = range a := -by rw [←bot_eq_zero, ←Iio_eq_Ico, Iio_eq_range] - -lemma _root_.finset.range_eq_Ico : range = Ico 0 := by { funext, exact (Ico_zero_eq_range n).symm } - -@[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := -by rw [Icc_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] - -@[simp] lemma card_Ico : (Ico a b).card = b - a := -by rw [Ico_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] - -@[simp] lemma card_Ioc : (Ioc a b).card = b - a := -by rw [Ioc_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] - -@[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := -by rw [Ioo_eq_range', list.card_to_finset, (list.nodup_range' _ _).erase_dup, list.length_range'] - -@[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := -by rw [←card_Icc, fintype.card_of_finset] - -@[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = b - a := -by rw [←card_Ico, fintype.card_of_finset] - -@[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = b - a := -by rw [←card_Ioc, fintype.card_of_finset] - -@[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := -by rw [←card_Ioo, fintype.card_of_finset] - --- TODO@Yaël: Generalize all the following lemmas to `succ_order` - -lemma Icc_succ_left : Icc a.succ b = Ioc a b := -by { ext x, rw [mem_Icc, mem_Ioc, succ_le_iff] } - -lemma Ico_succ_right : Ico a b.succ = Icc a b := -by { ext x, rw [mem_Ico, mem_Icc, lt_succ_iff] } - -lemma Ico_succ_left : Ico a.succ b = Ioo a b := -by { ext x, rw [mem_Ico, mem_Ioo, succ_le_iff] } - -lemma Icc_pred_right {b : ℕ} (h : 0 < b) : Icc a (b - 1) = Ico a b := -by { ext x, rw [mem_Icc, mem_Ico, lt_iff_le_pred h] } - -lemma Ico_succ_succ : Ico a.succ b.succ = Ioc a b := -by { ext x, rw [mem_Ico, mem_Ioc, succ_le_iff, lt_succ_iff] } - -@[simp] lemma Ico_succ_singleton : Ico a (a + 1) = {a} := -by rw [Ico_succ_right, Icc_self] - -@[simp] lemma Ico_pred_singleton {a : ℕ} (h : 0 < a) : Ico (a - 1) a = {a - 1} := -by rw [←Icc_pred_right _ h, Icc_self] - -variables {a b c} - -lemma Ico_succ_right_eq_insert_Ico (h : a ≤ b) : Ico a (b + 1) = insert b (Ico a b) := -by rw [Ico_succ_right, ←Ico_insert_right h] - -lemma Ico_insert_succ_left (h : a < b) : insert a (Ico a.succ b) = Ico a b := -by rw [Ico_succ_left, ←Ioo_insert_left h] - -lemma image_sub_const_Ico (h : c ≤ a) : - (Ico a b).image (λ x, x - c) = Ico (a - c) (b - c) := -begin - ext x, - rw mem_image, - split, - { rintro ⟨x, hx, rfl⟩, - rw [mem_Ico] at ⊢ hx, - exact ⟨sub_le_sub_right' hx.1 _, sub_lt_sub_right_of_le (h.trans hx.1) hx.2⟩ }, - { rintro h, - refine ⟨x + c, _, add_sub_cancel_right⟩, - rw mem_Ico at ⊢ h, - exact ⟨sub_le_iff_right.1 h.1, nat.add_lt_of_lt_sub_right h.2⟩ } -end - -lemma Ico_image_const_sub_eq_Ico (hac : a ≤ c) : - (Ico a b).image (λ x, c - x) = Ico (c + 1 - b) (c + 1 - a) := -begin - ext x, - rw [mem_image, mem_Ico], - split, - { rintro ⟨x, hx, rfl⟩, - rw mem_Ico at hx, - refine ⟨_, ((nat.sub_le_sub_left_iff hac).2 hx.1).trans_lt ((nat.sub_lt_sub_right_iff hac).2 - (nat.lt_succ_self _))⟩, - cases lt_or_le c b, - { rw nat.sub_eq_zero_of_le h, - exact zero_le _ }, - { rw ←succ_sub_succ c, - exact (nat.sub_le_sub_left_iff (succ_le_succ $ hx.2.le.trans h)).2 hx.2 } }, - { rintro ⟨hb, ha⟩, - rw [nat.lt_sub_left_iff_add_lt, lt_succ_iff] at ha, - have hx : x ≤ c := (nat.le_add_left _ _).trans ha, - refine ⟨c - x, _, nat.sub_sub_self hx⟩, - { rw mem_Ico, - refine ⟨nat.le_sub_right_of_add_le ha, _⟩, - rw [nat.sub_lt_left_iff_lt_add hx, ←succ_le_iff], - exact nat.le_add_of_sub_le_right hb } } -end - -lemma range_image_pred_top_sub (n : ℕ) : - (finset.range n).image (λ j, n - 1 - j) = finset.range n := -begin - cases n, - { rw [range_zero, image_empty] }, - { rw [finset.range_eq_Ico, Ico_image_const_sub_eq_Ico (zero_le _)], - simp_rw [succ_sub_succ, nat.sub_zero, nat.sub_self] } -end - -end nat diff --git a/src/data/pnat/interval.lean b/src/data/pnat/interval.lean index 3505c0fe5fff1..92090a77ac94e 100644 --- a/src/data/pnat/interval.lean +++ b/src/data/pnat/interval.lean @@ -20,12 +20,16 @@ namespace pnat variables (a b : ℕ+) lemma Icc_eq_finset_subtype : Icc a b = (Icc (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl +lemma Ico_eq_finset_subtype : Ico a b = (Ico (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl lemma map_subtype_embedding_Icc : (Icc a b).map (function.embedding.subtype _) = Icc (a : ℕ) b := map_subtype_embedding_Icc _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) +lemma map_subtype_embedding_Ico : (Ico a b).map (function.embedding.subtype _) = Ico (a : ℕ) b := +map_subtype_embedding_Ico _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) + lemma map_subtype_embedding_Ioc : (Ioc a b).map (function.embedding.subtype _) = Ioc (a : ℕ) b := map_subtype_embedding_Ioc _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) @@ -35,6 +39,9 @@ map_subtype_embedding_Ioo _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) @[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := by rw [←nat.card_Icc, ←map_subtype_embedding_Icc, card_map] +@[simp] lemma card_Ico : (Ico a b).card = b - a := +by rw [←nat.card_Ico, ←map_subtype_embedding_Ico, card_map] + @[simp] lemma card_Ioc : (Ioc a b).card = b - a := by rw [←nat.card_Ioc, ←map_subtype_embedding_Ioc, card_map] @@ -44,6 +51,9 @@ by rw [←nat.card_Ioo, ←map_subtype_embedding_Ioo, card_map] @[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := by rw [←card_Icc, fintype.card_of_finset] +@[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = b - a := +by rw [←card_Ico, fintype.card_of_finset] + @[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = b - a := by rw [←card_Ioc, fintype.card_of_finset] diff --git a/src/data/pnat/intervals.lean b/src/data/pnat/intervals.lean deleted file mode 100644 index ca6923a32d374..0000000000000 --- a/src/data/pnat/intervals.lean +++ /dev/null @@ -1,63 +0,0 @@ -/- -Copyright (c) 2020 Scott Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison, Yaël Dillies --/ -import data.nat.intervals - -/-! -# Finite intervals of positive naturals - -This file proves that `ℕ+` is a `locally_finite_order` and calculates the cardinality of its -intervals as finsets and fintypes. --/ - -open finset pnat - -instance : locally_finite_order ℕ+ := subtype.locally_finite_order _ - -namespace pnat -variables (a b : ℕ+) - -lemma Icc_eq_finset_subtype : Icc a b = (Icc (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl -lemma Ico_eq_finset_subtype : Ico a b = (Ico (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl -lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl -lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).subtype (λ (n : ℕ), 0 < n) := rfl - -lemma map_subtype_embedding_Icc : (Icc a b).map (function.embedding.subtype _) = Icc (a : ℕ) b := -map_subtype_embedding_Icc _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) - -lemma map_subtype_embedding_Ico : (Ico a b).map (function.embedding.subtype _) = Ico (a : ℕ) b := -map_subtype_embedding_Ico _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) - -lemma map_subtype_embedding_Ioc : (Ioc a b).map (function.embedding.subtype _) = Ioc (a : ℕ) b := -map_subtype_embedding_Ioc _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) - -lemma map_subtype_embedding_Ioo : (Ioo a b).map (function.embedding.subtype _) = Ioo (a : ℕ) b := -map_subtype_embedding_Ioo _ _ _ (λ c _ x hx _ hc _, hc.trans_le hx) - -@[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := -by rw [←nat.card_Icc, ←map_subtype_embedding_Icc, card_map] - -@[simp] lemma card_Ico : (Ico a b).card = b - a := -by rw [←nat.card_Ico, ←map_subtype_embedding_Ico, card_map] - -@[simp] lemma card_Ioc : (Ioc a b).card = b - a := -by rw [←nat.card_Ioc, ←map_subtype_embedding_Ioc, card_map] - -@[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := -by rw [←nat.card_Ioo, ←map_subtype_embedding_Ioo, card_map] - -@[simp] lemma card_fintype_Icc : fintype.card (set.Icc a b) = b + 1 - a := -by rw [←card_Icc, fintype.card_of_finset] - -@[simp] lemma card_fintype_Ico : fintype.card (set.Ico a b) = b - a := -by rw [←card_Ico, fintype.card_of_finset] - -@[simp] lemma card_fintype_Ioc : fintype.card (set.Ioc a b) = b - a := -by rw [←card_Ioc, fintype.card_of_finset] - -@[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := -by rw [←card_Ioo, fintype.card_of_finset] - -end pnat From f535f3ff194edb4e7d64c1e0fed57842bce6c8f1 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Wed, 6 Oct 2021 07:58:31 +0100 Subject: [PATCH 36/41] fix imports --- src/algebra/big_operators/intervals.lean | 2 +- src/data/finset/default.lean | 2 +- src/data/polynomial/inductions.lean | 2 +- src/data/polynomial/iterated_deriv.lean | 2 +- src/number_theory/ADE_inequality.lean | 2 +- src/number_theory/divisors.lean | 2 +- src/tactic/interval_cases.lean | 6 +++--- src/topology/metric_space/basic.lean | 2 +- src/topology/metric_space/emetric_space.lean | 2 +- test/fin_cases.lean | 2 +- 10 files changed, 12 insertions(+), 12 deletions(-) diff --git a/src/algebra/big_operators/intervals.lean b/src/algebra/big_operators/intervals.lean index b436e5317a249..d526afe54ffae 100644 --- a/src/algebra/big_operators/intervals.lean +++ b/src/algebra/big_operators/intervals.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl -/ import algebra.big_operators.basic -import data.nat.intervals +import data.nat.interval import tactic.linarith /-! diff --git a/src/data/finset/default.lean b/src/data/finset/default.lean index eb67d9d13773a..9450420fae5cf 100644 --- a/src/data/finset/default.lean +++ b/src/data/finset/default.lean @@ -1,6 +1,6 @@ import data.finset.basic import data.finset.fold -import data.finset.intervals +import data.finset.interval import data.finset.lattice import data.finset.nat_antidiagonal import data.finset.pi diff --git a/src/data/polynomial/inductions.lean b/src/data/polynomial/inductions.lean index aea91c1917668..44eee76787ed8 100644 --- a/src/data/polynomial/inductions.lean +++ b/src/data/polynomial/inductions.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Scott Morrison, Damiano Testa, Jens Wagemaker -/ -import data.nat.intervals +import data.nat.interval import data.polynomial.degree.definitions /-! diff --git a/src/data/polynomial/iterated_deriv.lean b/src/data/polynomial/iterated_deriv.lean index dae05263e7fb7..b17c7dc86a6bd 100644 --- a/src/data/polynomial/iterated_deriv.lean +++ b/src/data/polynomial/iterated_deriv.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ -import data.nat.intervals +import data.nat.interval import data.polynomial.derivative import tactic.linarith diff --git a/src/number_theory/ADE_inequality.lean b/src/number_theory/ADE_inequality.lean index 20913ffcd0fbe..a0e005f382dc4 100644 --- a/src/number_theory/ADE_inequality.lean +++ b/src/number_theory/ADE_inequality.lean @@ -5,7 +5,7 @@ Authors: Johan Commelin -/ import data.multiset.sort -import data.pnat.intervals +import data.pnat.interval import data.rat.order import tactic.norm_num diff --git a/src/number_theory/divisors.lean b/src/number_theory/divisors.lean index 807eab2902083..74836cc445837 100644 --- a/src/number_theory/divisors.lean +++ b/src/number_theory/divisors.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import algebra.big_operators.order -import data.nat.intervals +import data.nat.interval import data.nat.prime /-! diff --git a/src/tactic/interval_cases.lean b/src/tactic/interval_cases.lean index 5ed38a097166f..47b3a1811d6a1 100644 --- a/src/tactic/interval_cases.lean +++ b/src/tactic/interval_cases.lean @@ -3,9 +3,9 @@ Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import data.fin.intervals -import data.int.intervals -import data.pnat.intervals +import data.fin.interval +import data.int.interval +import data.pnat.interval import tactic.fin_cases /-! diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index a4bf601f2f3e7..7211eef887cda 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel -/ -import data.int.intervals +import data.int.interval import topology.algebra.ordered.basic import topology.metric_space.emetric_space diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index cf7c2673f56b5..025d39924564c 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -3,7 +3,7 @@ Copyright (c) 2015, 2017 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel -/ -import data.nat.intervals +import data.nat.interval import data.real.ennreal import topology.uniform_space.pi import topology.uniform_space.uniform_convergence diff --git a/test/fin_cases.lean b/test/fin_cases.lean index cf381372a47ed..9bea50de02267 100644 --- a/test/fin_cases.lean +++ b/test/fin_cases.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import data.nat.intervals +import data.nat.interval import data.nat.prime import group_theory.perm.sign import tactic.fin_cases From 7816e2ea58da2ad8777786f012f0137cdbcbced2 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Wed, 6 Oct 2021 09:18:36 +0100 Subject: [PATCH 37/41] merge data.finset.interval and data.finset.intervals --- src/data/finset/interval.lean | 152 +++++++++++++++++- src/data/finset/intervals.lean | 279 --------------------------------- 2 files changed, 149 insertions(+), 282 deletions(-) delete mode 100644 src/data/finset/intervals.lean diff --git a/src/data/finset/interval.lean b/src/data/finset/interval.lean index e1c4a14f2f505..4fa683f8e82d4 100644 --- a/src/data/finset/interval.lean +++ b/src/data/finset/interval.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2021 Yaël Dillies. All rights reserved. +Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yaël Dillies +Authors: Scott Morrison, Yaël Dillies -/ import order.locally_finite @@ -13,7 +13,9 @@ This file provides basic results about all the `finset.Ixx`, which are defined i ## TODO -Bring the lemmas about `finset.Ico` in `data.finset.intervals` here and in `data.nat.intervals`. +This file was originally only about `finset.Ico a b` where `a b : ℕ`. No care has yet been taken to +generalize these lemmas properly and many lemmas about `Icc`, `Ioc`, `Ioo` are missing. In general, +what's to do is taking the lemmas in `data.x.intervals` and abstract away the concrete structure. -/ namespace finset @@ -24,6 +26,9 @@ variables [preorder α] [locally_finite_order α] {a b : α} @[simp] lemma nonempty_Icc : (Icc a b).nonempty ↔ a ≤ b := by rw [←coe_nonempty, coe_Icc, set.nonempty_Icc] +@[simp] lemma nonempty_Ico : (Ico a b).nonempty ↔ a < b := +by rw [←coe_nonempty, coe_Ico, set.nonempty_Ico] + @[simp] lemma nonempty_Ioc : (Ioc a b).nonempty ↔ a < b := by rw [←coe_nonempty, coe_Ioc, set.nonempty_Ioc] @@ -33,6 +38,9 @@ by rw [←coe_nonempty, coe_Ioo, set.nonempty_Ioo] @[simp] lemma Icc_eq_empty_iff : Icc a b = ∅ ↔ ¬a ≤ b := by rw [←coe_eq_empty, coe_Icc, set.Icc_eq_empty_iff] +@[simp] lemma Ico_eq_empty_iff : Ico a b = ∅ ↔ ¬a < b := +by rw [←coe_eq_empty, coe_Ico, set.Ico_eq_empty_iff] + @[simp] lemma Ioc_eq_empty_iff : Ioc a b = ∅ ↔ ¬a < b := by rw [←coe_eq_empty, coe_Ioc, set.Ioc_eq_empty_iff] @@ -40,6 +48,7 @@ by rw [←coe_eq_empty, coe_Ioc, set.Ioc_eq_empty_iff] by rw [←coe_eq_empty, coe_Ioo, set.Ioo_eq_empty_iff] alias Icc_eq_empty_iff ↔ _ finset.Icc_eq_empty +alias Ico_eq_empty_iff ↔ _ finset.Ico_eq_empty alias Ioc_eq_empty_iff ↔ _ finset.Ioc_eq_empty @[simp] lemma Ioo_eq_empty (h : ¬a < b) : Ioo a b = ∅ := @@ -48,6 +57,9 @@ eq_empty_iff_forall_not_mem.2 $ λ x hx, h ((mem_Ioo.1 hx).1.trans (mem_Ioo.1 hx @[simp] lemma Icc_eq_empty_of_lt (h : b < a) : Icc a b = ∅ := Icc_eq_empty h.not_le +@[simp] lemma Ico_eq_empty_of_le (h : b ≤ a) : Ico a b = ∅ := +Ico_eq_empty h.not_lt + @[simp] lemma Ioc_eq_empty_of_le (h : b ≤ a) : Ioc a b = ∅ := Ioc_eq_empty h.not_lt @@ -56,12 +68,50 @@ Ioo_eq_empty h.not_lt variables (a) +@[simp] lemma Ico_self : Ico a a = ∅ := +by rw [←coe_eq_empty, coe_Ico, set.Ico_self] + @[simp] lemma Ioc_self : Ioc a a = ∅ := by rw [←coe_eq_empty, coe_Ioc, set.Ioc_self] @[simp] lemma Ioo_self : Ioo a a = ∅ := by rw [←coe_eq_empty, coe_Ioo, set.Ioo_self] +@[simp] lemma right_not_mem_Ico {a b : α} : b ∉ Ico a b := +by { rw [mem_Ico, not_and], exact λ _, lt_irrefl _ } + +lemma Ico_filter_lt_of_le_left [decidable_rel ((<) : α → α → Prop)] {a b c : α} (hca : c ≤ a) : + (Ico a b).filter (λ x, x < c) = ∅ := +finset.filter_false_of_mem (λ x hx, (hca.trans (mem_Ico.1 hx).1).not_lt) + +lemma Ico_filter_lt_of_right_le [decidable_rel ((<) : α → α → Prop)] {a b c : α} (hbc : b ≤ c) : + (Ico a b).filter (λ x, x < c) = Ico a b := +finset.filter_true_of_mem (λ x hx, (mem_Ico.1 hx).2.trans_le hbc) + +lemma Ico_filter_lt_of_le_right [decidable_rel ((<) : α → α → Prop)] {a b c : α} (hcb : c ≤ b) : + (Ico a b).filter (λ x, x < c) = Ico a c := +begin + ext x, + rw [mem_filter, mem_Ico, mem_Ico, and.right_comm], + exact and_iff_left_of_imp (λ h, h.2.trans_le hcb), +end + +lemma Ico_filter_le_of_le_left [decidable_rel ((≤) : α → α → Prop)] {a b c : α} (hca : c ≤ a) : + (Ico a b).filter (λ x, c ≤ x) = Ico a b := +finset.filter_true_of_mem (λ x hx, hca.trans (mem_Ico.1 hx).1) + +lemma Ico_filter_le_of_right_le [decidable_rel ((≤) : α → α → Prop)] {a b : α} : + (Ico a b).filter (λ x, b ≤ x) = ∅ := +finset.filter_false_of_mem (λ x hx, (mem_Ico.1 hx).2.not_le) + +lemma Ico_filter_le_of_left_le [decidable_rel ((≤) : α → α → Prop)] {a b c : α} (hac : a ≤ c) : + (Ico a b).filter (λ x, c ≤ x) = Ico c b := +begin + ext x, + rw [mem_filter, mem_Ico, mem_Ico, and_comm, and.left_comm], + exact and_iff_right_of_imp (λ h, hac.trans h.1), +end + end preorder section partial_order @@ -70,8 +120,88 @@ variables [partial_order α] [locally_finite_order α] {a b : α} @[simp] lemma Icc_self (a : α) : Icc a a = {a} := by rw [←coe_eq_singleton, coe_Icc, set.Icc_self] +lemma Ico_insert_right [decidable_eq α] (h : a ≤ b) : insert b (Ico a b) = Icc a b := +by rw [←coe_inj, coe_insert, coe_Icc, coe_Ico, set.insert_eq, set.union_comm, set.Ico_union_right h] + +lemma Ioo_insert_left [decidable_eq α] (h : a < b) : insert a (Ioo a b) = Ico a b := +by rw [←coe_inj, coe_insert, coe_Ioo, coe_Ico, set.insert_eq, set.union_comm, set.Ioo_union_left h] + +@[simp] lemma Ico_inter_Ico_consecutive [decidable_eq α] (a b c : α) : Ico a b ∩ Ico b c = ∅ := +begin + refine eq_empty_of_forall_not_mem (λ x hx, _), + rw [mem_inter, mem_Ico, mem_Ico] at hx, + exact hx.1.2.not_le hx.2.1, +end + +lemma Ico_disjoint_Ico_consecutive [decidable_eq α] (a b c : α) : disjoint (Ico a b) (Ico b c) := +le_of_eq $ Ico_inter_Ico_consecutive a b c + +lemma Ico_filter_le_left [decidable_rel ((≤) : α → α → Prop)] {a b : α} (hab : a < b) : + (Ico a b).filter (λ x, x ≤ a) = {a} := +begin + ext x, + rw [mem_filter, mem_Ico, mem_singleton, and.right_comm, ←le_antisymm_iff, eq_comm], + exact and_iff_left_of_imp (λ h, h.le.trans_lt hab), +end + end partial_order +section linear_order +variables [linear_order α] [locally_finite_order α] {a b : α} + +lemma Ico_subset_Ico_iff {a₁ b₁ a₂ b₂ : α} (h : a₁ < b₁) : + Ico a₁ b₁ ⊆ Ico a₂ b₂ ↔ a₂ ≤ a₁ ∧ b₁ ≤ b₂ := +by rw [←coe_subset, coe_Ico, coe_Ico, set.Ico_subset_Ico_iff h] + +lemma Ico_union_Ico_eq_Ico {a b c : α} (hab : a ≤ b) (hbc : b ≤ c) : + Ico a b ∪ Ico b c = Ico a c := +by rw [←coe_inj, coe_union, coe_Ico, coe_Ico, coe_Ico, set.Ico_union_Ico_eq_Ico hab hbc] + +lemma Ico_union_Ico' {a b c d : α} (hcb : c ≤ b) (had : a ≤ d) : + Ico a b ∪ Ico c d = Ico (min a c) (max b d) := +by rw [←coe_inj, coe_union, coe_Ico, coe_Ico, coe_Ico, set.Ico_union_Ico' hcb had] + +lemma Ico_union_Ico {a b c d : α} (h₁ : min a b ≤ max c d) (h₂ : min c d ≤ max a b) : + Ico a b ∪ Ico c d = Ico (min a c) (max b d) := +by rw [←coe_inj, coe_union, coe_Ico, coe_Ico, coe_Ico, set.Ico_union_Ico h₁ h₂] + +lemma Ico_inter_Ico {a b c d : α} : Ico a b ∩ Ico c d = Ico (max a c) (min b d) := +by rw [←coe_inj, coe_inter, coe_Ico, coe_Ico, coe_Ico, ←inf_eq_min, ←sup_eq_max, set.Ico_inter_Ico] + +@[simp] lemma Ico_filter_lt (a b c : α) : (Ico a b).filter (λ x, x < c) = Ico a (min b c) := +begin + cases le_total b c, + { rw [Ico_filter_lt_of_right_le h, min_eq_left h] }, + { rw [Ico_filter_lt_of_le_right h, min_eq_right h] } +end + +@[simp] lemma Ico_filter_le (a b c : α) : (Ico a b).filter (λ x, c ≤ x) = Ico (max a c) b := +begin + cases le_total a c, + { rw [Ico_filter_le_of_left_le h, max_eq_right h] }, + { rw [Ico_filter_le_of_le_left h, max_eq_left h] } +end + +@[simp] lemma Ico_diff_Ico_left (a b c : α) : (Ico a b) \ (Ico a c) = Ico (max a c) b := +begin + cases le_total a c, + { ext x, + rw [mem_sdiff, mem_Ico, mem_Ico, mem_Ico, max_eq_right h, and.right_comm, not_and, not_lt], + exact and_congr_left' ⟨λ hx, hx.2 hx.1, λ hx, ⟨h.trans hx, λ _, hx⟩⟩ }, + { rw [Ico_eq_empty_of_le h, sdiff_empty, max_eq_left h] } +end + +@[simp] lemma Ico_diff_Ico_right (a b c : α) : (Ico a b) \ (Ico c b) = Ico a (min b c) := +begin + cases le_total b c, + { rw [Ico_eq_empty_of_le h, sdiff_empty, min_eq_left h] }, + { ext x, + rw [mem_sdiff, mem_Ico, mem_Ico, mem_Ico, min_eq_right h, and_assoc, not_and', not_le], + exact and_congr_right' ⟨λ hx, hx.2 hx.1, λ hx, ⟨hx.trans_le h, λ _, hx⟩⟩ } +end + +end linear_order + section ordered_cancel_add_comm_monoid variables [ordered_cancel_add_comm_monoid α] [has_exists_add_of_le α] [decidable_eq α] [locally_finite_order α] @@ -92,6 +222,22 @@ begin exact ⟨a + y, mem_Icc.2 ⟨le_of_add_le_add_right hx.1, le_of_add_le_add_right hx.2⟩, hy⟩ } end +lemma image_add_const_Ico (a b c : α) : (Ico a b).image ((+) c) = Ico (a + c) (b + c) := +begin + ext x, + rw [mem_image, mem_Ico], + split, + { rintro ⟨y, hy, hx⟩, + rw mem_Ico at hy, + rw [←hx, add_comm c], + exact ⟨add_le_add_right hy.1 c, add_lt_add_right hy.2 c⟩ }, + { intro hx, + obtain ⟨y, hy⟩ := exists_add_of_le hx.1, + rw [hy, add_right_comm] at hx, + rw [eq_comm, add_right_comm, add_comm] at hy, + exact ⟨a + y, mem_Ico.2 ⟨le_of_add_le_add_right hx.1, lt_of_add_lt_add_right hx.2⟩, hy⟩ } +end + lemma image_add_const_Ioc (a b c : α) : (Ioc a b).image ((+) c) = Ioc (a + c) (b + c) := begin ext x, diff --git a/src/data/finset/intervals.lean b/src/data/finset/intervals.lean deleted file mode 100644 index 86133763f78f0..0000000000000 --- a/src/data/finset/intervals.lean +++ /dev/null @@ -1,279 +0,0 @@ -/- -Copyright (c) 2019 Scott Morrison. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Scott Morrison, Yaël Dillies --/ -import order.locally_finite - -/-! -# Intervals as finsets - -This file provides basic results about all the `finset.Ixx`, which are defined in -`order.locally_finite`. - -## TODO - -This file was originally only about `finset.Ico a b` where `a b : ℕ`. No care has yet been taken to -generalize these lemmas properly and many lemmas about `Icc`, `Ioc`, `Ioo` are missing. In general, -what's to do is taking the lemmas in `data.x.intervals` and abstract away the concrete structure. --/ - -lemma finset.coe_eq_singleton {α : Type*} {s : finset α} {a : α} : - (s : set α) = {a} ↔ s = {a} := -by rw [←finset.coe_singleton, finset.coe_inj] - -namespace finset -variables {α : Type*} -section preorder -variables [preorder α] [locally_finite_order α] {a b : α} - -@[simp] lemma nonempty_Icc : (Icc a b).nonempty ↔ a ≤ b := -by rw [←coe_nonempty, coe_Icc, set.nonempty_Icc] - -@[simp] lemma nonempty_Ico : (Ico a b).nonempty ↔ a < b := -by rw [←coe_nonempty, coe_Ico, set.nonempty_Ico] - -@[simp] lemma nonempty_Ioc : (Ioc a b).nonempty ↔ a < b := -by rw [←coe_nonempty, coe_Ioc, set.nonempty_Ioc] - -@[simp] lemma nonempty_Ioo [densely_ordered α] : (Ioo a b).nonempty ↔ a < b := -by rw [←coe_nonempty, coe_Ioo, set.nonempty_Ioo] - -@[simp] lemma Icc_eq_empty_iff : Icc a b = ∅ ↔ ¬a ≤ b := -by rw [←coe_eq_empty, coe_Icc, set.Icc_eq_empty_iff] - -@[simp] lemma Ico_eq_empty_iff : Ico a b = ∅ ↔ ¬a < b := -by rw [←coe_eq_empty, coe_Ico, set.Ico_eq_empty_iff] - -@[simp] lemma Ioc_eq_empty_iff : Ioc a b = ∅ ↔ ¬a < b := -by rw [←coe_eq_empty, coe_Ioc, set.Ioc_eq_empty_iff] - -@[simp] lemma Ioo_eq_empty_iff [densely_ordered α] : Ioo a b = ∅ ↔ ¬a < b := -by rw [←coe_eq_empty, coe_Ioo, set.Ioo_eq_empty_iff] - -alias Icc_eq_empty_iff ↔ _ finset.Icc_eq_empty -alias Ico_eq_empty_iff ↔ _ finset.Ico_eq_empty -alias Ioc_eq_empty_iff ↔ _ finset.Ioc_eq_empty - -@[simp] lemma Ioo_eq_empty (h : ¬a < b) : Ioo a b = ∅ := -eq_empty_iff_forall_not_mem.2 $ λ x hx, h ((mem_Ioo.1 hx).1.trans (mem_Ioo.1 hx).2) - -@[simp] lemma Icc_eq_empty_of_lt (h : b < a) : Icc a b = ∅ := -Icc_eq_empty h.not_le - -@[simp] lemma Ico_eq_empty_of_le (h : b ≤ a) : Ico a b = ∅ := -Ico_eq_empty h.not_lt - -@[simp] lemma Ioc_eq_empty_of_le (h : b ≤ a) : Ioc a b = ∅ := -Ioc_eq_empty h.not_lt - -@[simp] lemma Ioo_eq_empty_of_le (h : b ≤ a) : Ioo a b = ∅ := -Ioo_eq_empty h.not_lt - -variables (a) - -@[simp] lemma Ico_self : Ico a a = ∅ := -by rw [←coe_eq_empty, coe_Ico, set.Ico_self] - -@[simp] lemma Ioc_self : Ioc a a = ∅ := -by rw [←coe_eq_empty, coe_Ioc, set.Ioc_self] - -@[simp] lemma Ioo_self : Ioo a a = ∅ := -by rw [←coe_eq_empty, coe_Ioo, set.Ioo_self] - -@[simp] lemma right_not_mem_Ico {a b : α} : b ∉ Ico a b := -by { rw [mem_Ico, not_and], exact λ _, lt_irrefl _ } - -lemma Ico_filter_lt_of_le_left [decidable_rel ((<) : α → α → Prop)] {a b c : α} (hca : c ≤ a) : - (Ico a b).filter (λ x, x < c) = ∅ := -finset.filter_false_of_mem (λ x hx, (hca.trans (mem_Ico.1 hx).1).not_lt) - -lemma Ico_filter_lt_of_right_le [decidable_rel ((<) : α → α → Prop)] {a b c : α} (hbc : b ≤ c) : - (Ico a b).filter (λ x, x < c) = Ico a b := -finset.filter_true_of_mem (λ x hx, (mem_Ico.1 hx).2.trans_le hbc) - -lemma Ico_filter_lt_of_le_right [decidable_rel ((<) : α → α → Prop)] {a b c : α} (hcb : c ≤ b) : - (Ico a b).filter (λ x, x < c) = Ico a c := -begin - ext x, - rw [mem_filter, mem_Ico, mem_Ico, and.right_comm], - exact and_iff_left_of_imp (λ h, h.2.trans_le hcb), -end - -lemma Ico_filter_le_of_le_left [decidable_rel ((≤) : α → α → Prop)] {a b c : α} (hca : c ≤ a) : - (Ico a b).filter (λ x, c ≤ x) = Ico a b := -finset.filter_true_of_mem (λ x hx, hca.trans (mem_Ico.1 hx).1) - -lemma Ico_filter_le_of_right_le [decidable_rel ((≤) : α → α → Prop)] {a b : α} : - (Ico a b).filter (λ x, b ≤ x) = ∅ := -finset.filter_false_of_mem (λ x hx, (mem_Ico.1 hx).2.not_le) - -lemma Ico_filter_le_of_left_le [decidable_rel ((≤) : α → α → Prop)] {a b c : α} (hac : a ≤ c) : - (Ico a b).filter (λ x, c ≤ x) = Ico c b := -begin - ext x, - rw [mem_filter, mem_Ico, mem_Ico, and_comm, and.left_comm], - exact and_iff_right_of_imp (λ h, hac.trans h.1), -end - -end preorder - -section partial_order -variables [partial_order α] [locally_finite_order α] {a b : α} - - -@[simp] lemma Icc_self (a : α) : Icc a a = {a} := -by rw [←coe_eq_singleton, coe_Icc, set.Icc_self] - -lemma Ico_insert_right [decidable_eq α] (h : a ≤ b) : insert b (Ico a b) = Icc a b := -by rw [←coe_inj, coe_insert, coe_Icc, coe_Ico, set.insert_eq, set.union_comm, set.Ico_union_right h] - -lemma Ioo_insert_left [decidable_eq α] (h : a < b) : insert a (Ioo a b) = Ico a b := -by rw [←coe_inj, coe_insert, coe_Ioo, coe_Ico, set.insert_eq, set.union_comm, set.Ioo_union_left h] - -@[simp] lemma Ico_inter_Ico_consecutive [decidable_eq α] (a b c : α) : Ico a b ∩ Ico b c = ∅ := -begin - refine eq_empty_of_forall_not_mem (λ x hx, _), - rw [mem_inter, mem_Ico, mem_Ico] at hx, - exact hx.1.2.not_le hx.2.1, -end - -lemma Ico_disjoint_Ico_consecutive [decidable_eq α] (a b c : α) : disjoint (Ico a b) (Ico b c) := -le_of_eq $ Ico_inter_Ico_consecutive a b c - -lemma Ico_filter_le_left [decidable_rel ((≤) : α → α → Prop)] {a b : α} (hab : a < b) : - (Ico a b).filter (λ x, x ≤ a) = {a} := -begin - ext x, - rw [mem_filter, mem_Ico, mem_singleton, and.right_comm, ←le_antisymm_iff, eq_comm], - exact and_iff_left_of_imp (λ h, h.le.trans_lt hab), -end - -end partial_order - -section linear_order -variables [linear_order α] [locally_finite_order α] {a b : α} - -lemma Ico_subset_Ico_iff {a₁ b₁ a₂ b₂ : α} (h : a₁ < b₁) : - Ico a₁ b₁ ⊆ Ico a₂ b₂ ↔ a₂ ≤ a₁ ∧ b₁ ≤ b₂ := -by rw [←coe_subset, coe_Ico, coe_Ico, set.Ico_subset_Ico_iff h] - -lemma Ico_union_Ico_eq_Ico {a b c : α} (hab : a ≤ b) (hbc : b ≤ c) : - Ico a b ∪ Ico b c = Ico a c := -by rw [←coe_inj, coe_union, coe_Ico, coe_Ico, coe_Ico, set.Ico_union_Ico_eq_Ico hab hbc] - -lemma Ico_union_Ico' {a b c d : α} (hcb : c ≤ b) (had : a ≤ d) : - Ico a b ∪ Ico c d = Ico (min a c) (max b d) := -by rw [←coe_inj, coe_union, coe_Ico, coe_Ico, coe_Ico, set.Ico_union_Ico' hcb had] - -lemma Ico_union_Ico {a b c d : α} (h₁ : min a b ≤ max c d) (h₂ : min c d ≤ max a b) : - Ico a b ∪ Ico c d = Ico (min a c) (max b d) := -by rw [←coe_inj, coe_union, coe_Ico, coe_Ico, coe_Ico, set.Ico_union_Ico h₁ h₂] - -lemma Ico_inter_Ico {a b c d : α} : Ico a b ∩ Ico c d = Ico (max a c) (min b d) := -by rw [←coe_inj, coe_inter, coe_Ico, coe_Ico, coe_Ico, ←inf_eq_min, ←sup_eq_max, set.Ico_inter_Ico] - -@[simp] lemma Ico_filter_lt (a b c : α) : (Ico a b).filter (λ x, x < c) = Ico a (min b c) := -begin - cases le_total b c, - { rw [Ico_filter_lt_of_right_le h, min_eq_left h] }, - { rw [Ico_filter_lt_of_le_right h, min_eq_right h] } -end - -@[simp] lemma Ico_filter_le (a b c : α) : (Ico a b).filter (λ x, c ≤ x) = Ico (max a c) b := -begin - cases le_total a c, - { rw [Ico_filter_le_of_left_le h, max_eq_right h] }, - { rw [Ico_filter_le_of_le_left h, max_eq_left h] } -end - -@[simp] lemma Ico_diff_Ico_left (a b c : α) : (Ico a b) \ (Ico a c) = Ico (max a c) b := -begin - cases le_total a c, - { ext x, - rw [mem_sdiff, mem_Ico, mem_Ico, mem_Ico, max_eq_right h, and.right_comm, not_and, not_lt], - exact and_congr_left' ⟨λ hx, hx.2 hx.1, λ hx, ⟨h.trans hx, λ _, hx⟩⟩ }, - { rw [Ico_eq_empty_of_le h, sdiff_empty, max_eq_left h] } -end - -@[simp] lemma Ico_diff_Ico_right (a b c : α) : (Ico a b) \ (Ico c b) = Ico a (min b c) := -begin - cases le_total b c, - { rw [Ico_eq_empty_of_le h, sdiff_empty, min_eq_left h] }, - { ext x, - rw [mem_sdiff, mem_Ico, mem_Ico, mem_Ico, min_eq_right h, and_assoc, not_and', not_le], - exact and_congr_right' ⟨λ hx, hx.2 hx.1, λ hx, ⟨hx.trans_le h, λ _, hx⟩⟩ } -end - -end linear_order - -section ordered_cancel_add_comm_monoid -variables [ordered_cancel_add_comm_monoid α] [has_exists_add_of_le α] [decidable_eq α] - [locally_finite_order α] - -lemma image_add_const_Icc (a b c : α) : (Icc a b).image ((+) c) = Icc (a + c) (b + c) := -begin - ext x, - rw [mem_image, mem_Icc], - split, - { rintro ⟨y, hy, rfl⟩, - rw mem_Icc at hy, - rw add_comm c, - exact ⟨add_le_add_right hy.1 c, add_le_add_right hy.2 c⟩ }, - { intro hx, - obtain ⟨y, hy⟩ := exists_add_of_le hx.1, - rw [hy, add_right_comm] at hx, - rw [eq_comm, add_right_comm, add_comm] at hy, - exact ⟨a + y, mem_Icc.2 ⟨le_of_add_le_add_right hx.1, le_of_add_le_add_right hx.2⟩, hy⟩ } -end - -lemma image_add_const_Ico (a b c : α) : (Ico a b).image ((+) c) = Ico (a + c) (b + c) := -begin - ext x, - rw [mem_image, mem_Ico], - split, - { rintro ⟨y, hy, hx⟩, - rw mem_Ico at hy, - rw [←hx, add_comm c], - exact ⟨add_le_add_right hy.1 c, add_lt_add_right hy.2 c⟩ }, - { intro hx, - obtain ⟨y, hy⟩ := exists_add_of_le hx.1, - rw [hy, add_right_comm] at hx, - rw [eq_comm, add_right_comm, add_comm] at hy, - exact ⟨a + y, mem_Ico.2 ⟨le_of_add_le_add_right hx.1, lt_of_add_lt_add_right hx.2⟩, hy⟩ } -end - -lemma image_add_const_Ioc (a b c : α) : (Ioc a b).image ((+) c) = Ioc (a + c) (b + c) := -begin - ext x, - rw [mem_image, mem_Ioc], - split, - { rintro ⟨y, hy, rfl⟩, - rw mem_Ioc at hy, - rw add_comm c, - exact ⟨add_lt_add_right hy.1 c, add_le_add_right hy.2 c⟩ }, - { intro hx, - obtain ⟨y, hy⟩ := exists_add_of_le hx.1.le, - rw [hy, add_right_comm] at hx, - rw [eq_comm, add_right_comm, add_comm] at hy, - exact ⟨a + y, mem_Ioc.2 ⟨lt_of_add_lt_add_right hx.1, le_of_add_le_add_right hx.2⟩, hy⟩ } -end - -lemma image_add_const_Ioo (a b c : α) : (Ioo a b).image ((+) c) = Ioo (a + c) (b + c) := -begin - ext x, - rw [mem_image, mem_Ioo], - split, - { rintro ⟨y, hy, rfl⟩, - rw mem_Ioo at hy, - rw add_comm c, - exact ⟨add_lt_add_right hy.1 c, add_lt_add_right hy.2 c⟩ }, - { intro hx, - obtain ⟨y, hy⟩ := exists_add_of_le hx.1.le, - rw [hy, add_right_comm] at hx, - rw [eq_comm, add_right_comm, add_comm] at hy, - exact ⟨a + y, mem_Ioo.2 ⟨lt_of_add_lt_add_right hx.1, lt_of_add_lt_add_right hx.2⟩, hy⟩ } -end - -end ordered_cancel_add_comm_monoid -end finset From ec30f389944c7ac3c327083e1fa5c1c3a5aa9518 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Wed, 6 Oct 2021 11:03:09 +0100 Subject: [PATCH 38/41] fix data.nat.interval --- src/data/nat/interval.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/nat/interval.lean b/src/data/nat/interval.lean index 492a8b82f14bd..9a91d96cf99e6 100644 --- a/src/data/nat/interval.lean +++ b/src/data/nat/interval.lean @@ -134,7 +134,7 @@ begin rw [mem_Ico] at ⊢ hx, exact ⟨sub_le_sub_right' hx.1 _, sub_lt_sub_right_of_le (h.trans hx.1) hx.2⟩ }, { rintro h, - refine ⟨x + c, _, add_sub_cancel_right⟩, + refine ⟨x + c, _, add_sub_cancel_right _ _⟩, rw mem_Ico at ⊢ h, exact ⟨sub_le_iff_right.1 h.1, nat.add_lt_of_lt_sub_right h.2⟩ } end From 5e64c103c56a200a1f43ee731a7e43ce94f852f5 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Wed, 6 Oct 2021 19:53:03 +0100 Subject: [PATCH 39/41] update ring_theory.henselian --- src/ring_theory/henselian.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/ring_theory/henselian.lean b/src/ring_theory/henselian.lean index 4136ed06e7e2e..b951a6860bb8f 100644 --- a/src/ring_theory/henselian.lean +++ b/src/ring_theory/henselian.lean @@ -219,7 +219,7 @@ instance is_adic_complete.henselian_ring mul_neg_eq_neg_mul_symm, finset.sum_singleton, finset.range_one, pow_zero], rw [mul_left_comm, ring.mul_inverse_cancel _ (hf'c n), mul_one, add_neg_self], exact ideal.zero_mem _ }, - { refine submodule.sum_mem _ _, simp only [finset.Ico.mem], + { refine submodule.sum_mem _ _, simp only [finset.mem_Ico], rintro i ⟨h2i, hi⟩, have aux : n + 2 ≤ i * (n + 1), { transitivity 2 * (n + 1); nlinarith only [h2i], }, refine ideal.mul_mem_left _ _ (ideal.pow_le_pow aux _), From 9bb04039bd805f99ebb874386573be397d7029da Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Wed, 6 Oct 2021 20:23:41 +0100 Subject: [PATCH 40/41] bump --- src/data/nat/interval.lean | 13 ++++++------- 1 file changed, 6 insertions(+), 7 deletions(-) diff --git a/src/data/nat/interval.lean b/src/data/nat/interval.lean index 4ee8a911e0824..c24e63727606b 100644 --- a/src/data/nat/interval.lean +++ b/src/data/nat/interval.lean @@ -137,7 +137,7 @@ begin { rintro h, refine ⟨x + c, _, add_sub_cancel_right _ _⟩, rw mem_Ico at ⊢ h, - exact ⟨sub_le_iff_right.1 h.1, nat.add_lt_of_lt_sub_right h.2⟩ } + exact ⟨sub_le_iff_right.1 h.1, lt_sub_iff_right.1 h.2⟩ } end lemma Ico_image_const_sub_eq_Ico (hac : a ≤ c) : @@ -148,21 +148,20 @@ begin split, { rintro ⟨x, hx, rfl⟩, rw mem_Ico at hx, - refine ⟨_, ((nat.sub_le_sub_left_iff hac).2 hx.1).trans_lt ((nat.sub_lt_sub_right_iff hac).2 + refine ⟨_, ((sub_le_sub_iff_left' hac).2 hx.1).trans_lt ((sub_lt_sub_iff_right' hac).2 (nat.lt_succ_self _))⟩, cases lt_or_le c b, { rw nat.sub_eq_zero_of_le h, exact zero_le _ }, { rw ←succ_sub_succ c, - exact (nat.sub_le_sub_left_iff (succ_le_succ $ hx.2.le.trans h)).2 hx.2 } }, + exact (sub_le_sub_iff_left' (succ_le_succ $ hx.2.le.trans h)).2 hx.2 } }, { rintro ⟨hb, ha⟩, - rw [nat.lt_sub_left_iff_add_lt, lt_succ_iff] at ha, + rw [lt_sub_iff_left, lt_succ_iff] at ha, have hx : x ≤ c := (nat.le_add_left _ _).trans ha, refine ⟨c - x, _, nat.sub_sub_self hx⟩, { rw mem_Ico, - refine ⟨nat.le_sub_right_of_add_le ha, _⟩, - rw [nat.sub_lt_left_iff_lt_add hx, ←succ_le_iff], - exact nat.le_add_of_sub_le_right hb } } + refine ⟨le_sub_of_add_le_right' ha, _⟩, + rwa [sub_lt_iff_left hx, ←succ_le_iff, sub_le_iff_right] } } end lemma range_image_pred_top_sub (n : ℕ) : From f48e637e6553ab85578590929303b42532cc53c0 Mon Sep 17 00:00:00 2001 From: YaelDillies Date: Wed, 6 Oct 2021 22:24:04 +0100 Subject: [PATCH 41/41] fix data.nat.interval --- src/data/nat/interval.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/data/nat/interval.lean b/src/data/nat/interval.lean index c24e63727606b..c6cd7876e29a9 100644 --- a/src/data/nat/interval.lean +++ b/src/data/nat/interval.lean @@ -160,8 +160,8 @@ begin have hx : x ≤ c := (nat.le_add_left _ _).trans ha, refine ⟨c - x, _, nat.sub_sub_self hx⟩, { rw mem_Ico, - refine ⟨le_sub_of_add_le_right' ha, _⟩, - rwa [sub_lt_iff_left hx, ←succ_le_iff, sub_le_iff_right] } } + exact ⟨le_sub_of_add_le_right' ha, (sub_lt_iff_left hx).2 $ succ_le_iff.1 $ + sub_le_iff_right.1 hb⟩ } } end lemma range_image_pred_top_sub (n : ℕ) :