Skip to content

Commit

Permalink
feat(data/finset,order/conditionally_complete_lattice): lemmas about …
Browse files Browse the repository at this point in the history
…`min'/max'` (#8782)

## `data/finset/*`

* add `finset.nonempty.to_set`;
* add lemmas `finset.max'_lt_iff`, `finset.lt_min'_iff`,
  `finset.max'_eq_sup'`, `finset.min'_eq_inf'`;
* rewrite `finset.induction_on_max` without using `finset.card`,
  move one step to `finset.lt_max'_of_mem_erase_max'`;

## `order/conditionally_complete_lattice`

* add lemmas relating `Sup`/`Inf` of a nonempty finite set in a
  conditionally complete lattice to
 `finset.sup'`/`finset.inf'`/`finset.max'`/`finset.min'`;
* a few more lemmas about `Sup`/`Inf` of a nonempty finite set
  in a conditionally complete lattice / linear order;

## `order/filter/at_top_bot`

* golf the proof of `filter.high_scores`.
  • Loading branch information
urkud committed Aug 22, 2021
1 parent ea9cd02 commit db9d4a3
Show file tree
Hide file tree
Showing 4 changed files with 63 additions and 48 deletions.
2 changes: 2 additions & 0 deletions src/data/finset/basic.lean
Expand Up @@ -304,6 +304,8 @@ protected def nonempty (s : finset α) : Prop := ∃ x:α, x ∈ s

@[simp, norm_cast] lemma coe_nonempty {s : finset α} : (s:set α).nonempty ↔ s.nonempty := iff.rfl

alias coe_nonempty ↔ _ finset.nonempty.to_set

lemma nonempty.bex {s : finset α} (h : s.nonempty) : ∃ x:α, x ∈ s := h

lemma nonempty.mono {s t : finset α} (hst : s ⊆ t) (hs : s.nonempty) : t.nonempty :=
Expand Down
34 changes: 26 additions & 8 deletions src/data/finset/lattice.lean
Expand Up @@ -656,6 +656,18 @@ theorem is_greatest_max' : is_greatest ↑s (s.max' H) := ⟨max'_mem _ _, le_ma
@[simp] lemma max'_le_iff {x} : s.max' H ≤ x ↔ ∀ y ∈ s, y ≤ x :=
is_lub_le_iff (is_greatest_max' s H).is_lub

@[simp] lemma max'_lt_iff {x} : s.max' H < x ↔ ∀ y ∈ s, y < x :=
⟨λ Hlt y hy, (s.le_max' y hy).trans_lt Hlt, λ H, H _ $ s.max'_mem _⟩

@[simp] lemma lt_min'_iff {x} : x < s.min' H ↔ ∀ y ∈ s, x < y :=
@max'_lt_iff (order_dual α) _ _ H _

lemma max'_eq_sup' : s.max' H = s.sup' H id :=
eq_of_forall_ge_iff $ λ a, (max'_le_iff _ _).trans (sup'_le_iff _ _).symm

lemma min'_eq_inf' : s.min' H = s.inf' H id :=
@max'_eq_sup' (order_dual α) _ s H

/-- `{a}.max' _` is `a`. -/
@[simp] lemma max'_singleton (a : α) :
({a} : finset α).max' (singleton_nonempty _) = a :=
Expand Down Expand Up @@ -717,6 +729,14 @@ lemma min'_insert (a : α) (s : finset α) (H : s.nonempty) :
(is_least_min' _ _).unique $
by { rw [coe_insert, min_comm], exact (is_least_min' _ _).insert _ }

lemma lt_max'_of_mem_erase_max' [decidable_eq α] {a : α} (ha : a ∈ s.erase (s.max' H)) :
a < s.max' H :=
lt_of_le_of_ne (le_max' _ _ (mem_of_mem_erase ha)) $ ne_of_mem_of_not_mem ha $ not_mem_erase _ _

lemma min'_lt_of_mem_erase_min' [decidable_eq α] {a : α} (ha : a ∈ s.erase (s.min' H)) :
s.min' H < a :=
@lt_max'_of_mem_erase_max' (order_dual α) _ s H _ a ha

/-- Induction principle for `finset`s in a linearly ordered type: a predicate is true on all
`s : finset α` provided that:
Expand All @@ -727,14 +747,12 @@ lemma min'_insert (a : α) (s : finset α) (H : s.nonempty) :
lemma induction_on_max [decidable_eq α] {p : finset α → Prop} (s : finset α) (h0 : p ∅)
(step : ∀ a s, (∀ x ∈ s, x < a) → p s → p (insert a s)) : p s :=
begin
induction hn : s.card with n ihn generalizing s,
{ rwa [card_eq_zero.1 hn] },
{ have A : s.nonempty, from card_pos.1 (hn.symm ▸ n.succ_pos),
have B : s.max' A ∈ s, from max'_mem s A,
rw [← insert_erase B],
refine step _ _ (λ x hx, _) (ihn _ _),
{ rw [mem_erase] at hx, exact (le_max' s x hx.2).lt_of_ne hx.1 },
{ rw [card_erase_of_mem B, hn, nat.pred_succ] } }
induction s using finset.strong_induction_on with s ihs,
rcases s.eq_empty_or_nonempty with rfl|hne,
{ exact h0 },
{ have H : s.max' hne ∈ s, from max'_mem s hne,
rw ← insert_erase H,
exact step _ _ (λ x, s.lt_max'_of_mem_erase_max' hne) (ihs _ $ erase_ssubset H) }
end

/-- Induction principle for `finset`s in a linearly ordered type: a predicate is true on all
Expand Down
43 changes: 26 additions & 17 deletions src/order/conditionally_complete_lattice.lean
Expand Up @@ -473,6 +473,15 @@ lemma csupr_mem_Inter_Icc_of_mono_decr_Icc_nat
csupr_mem_Inter_Icc_of_mono_decr_Icc
(@monotone_nat_of_le_succ (order_dual $ set α) _ (λ n, Icc (f n) (g n)) h) h'

lemma finset.nonempty.sup'_eq_cSup_image {s : finset β} (hs : s.nonempty) (f : β → α) :
s.sup' hs f = Sup (f '' s) :=
eq_of_forall_ge_iff $ λ a,
by simp [cSup_le_iff (s.finite_to_set.image f).bdd_above (hs.to_set.image f)]

lemma finset.nonempty.sup'_id_eq_cSup {s : finset α} (hs : s.nonempty) :
s.sup' hs id = Sup s :=
by rw [hs.sup'_eq_cSup_image, image_id]

end conditionally_complete_lattice

instance pi.conditionally_complete_lattice {ι : Type*} {α : Π i : ι, Type*}
Expand All @@ -491,29 +500,29 @@ instance pi.conditionally_complete_lattice {ι : Type*} {α : Π i : ι, Type*}
section conditionally_complete_linear_order
variables [conditionally_complete_linear_order α] {s t : set α} {a b : α}

lemma set.nonempty.cSup_mem (h : s.nonempty) (hs : finite s) : Sup s ∈ s :=
begin
classical,
revert h,
apply finite.induction_on hs,
{ simp },
rintros a t hat t_fin ih -,
rcases t.eq_empty_or_nonempty with rfl | ht,
{ simp },
{ rw cSup_insert t_fin.bdd_above ht,
by_cases ha : a ≤ Sup t,
{ simp [sup_eq_right.mpr ha, ih ht] },
{ simp only [sup_eq_left, mem_insert_iff, (not_le.mp ha).le, true_or] } }
end
lemma finset.nonempty.cSup_eq_max' {s : finset α} (h : s.nonempty) : Sup ↑s = s.max' h :=
eq_of_forall_ge_iff $ λ a, (cSup_le_iff s.bdd_above h.to_set).trans (s.max'_le_iff h).symm

lemma finset.nonempty.cInf_eq_min' {s : finset α} (h : s.nonempty) : Inf ↑s = s.min' h :=
@finset.nonempty.cSup_eq_max' (order_dual α) _ s h

lemma finset.nonempty.cSup_mem {s : finset α} (h : s.nonempty) : Sup (s : set α) ∈ s :=
set.nonempty.cSup_mem h s.finite_to_set
by { rw h.cSup_eq_max', exact s.max'_mem _ }

lemma finset.nonempty.cInf_mem {s : finset α} (h : s.nonempty) : Inf (s : set α) ∈ s :=
@finset.nonempty.cSup_mem (order_dual α) _ _ h

lemma set.nonempty.cSup_mem (h : s.nonempty) (hs : finite s) : Sup s ∈ s :=
by { unfreezingI { lift s to finset α using hs }, exact finset.nonempty.cSup_mem h }

lemma set.nonempty.cInf_mem (h : s.nonempty) (hs : finite s) : Inf s ∈ s :=
@set.nonempty.cSup_mem (order_dual α) _ _ h hs

lemma finset.nonempty.cInf_mem {s : finset α} (h : s.nonempty) : Inf (s : set α) ∈ s :=
set.nonempty.cInf_mem h s.finite_to_set
lemma set.finite.cSup_lt_iff (hs : finite s) (h : s.nonempty) : Sup s < a ↔ ∀ x ∈ s, x < a :=
⟨λ h x hx, (le_cSup hs.bdd_above hx).trans_lt h, λ H, H _ $ h.cSup_mem hs⟩

lemma set.finite.lt_cInf_iff (hs : finite s) (h : s.nonempty) : a < Inf s ↔ ∀ x ∈ s, a < x :=
@set.finite.cSup_lt_iff (order_dual α) _ _ _ hs h

/-- When b < Sup s, there is an element a in s with b < a, if s is nonempty and the order is
a linear order. -/
Expand Down
32 changes: 9 additions & 23 deletions src/order/filter/at_top_bot.lean
Expand Up @@ -310,30 +310,16 @@ lemma high_scores [linear_order β] [no_top_order β] {u : ℕ → β}
(hu : tendsto u at_top at_top) : ∀ N, ∃ n ≥ N, ∀ k < n, u k < u n :=
begin
intros N,
let A := finset.image u (finset.range $ N+1), -- A = {u 0, ..., u N}
have Ane : A.nonempty,
from ⟨u 0, finset.mem_image_of_mem _ (finset.mem_range.mpr $ nat.zero_lt_succ _)⟩,
let M := finset.max' A Ane,
have ex : ∃ n ≥ N, M < u n,
obtain ⟨k, hkn, hku⟩ : ∃ k ≤ N, ∀ l ≤ N, u l ≤ u k,
from exists_max_image _ u (finite_le_nat N) ⟨N, le_refl N⟩,
have ex : ∃ n ≥ N, u k < u n,
from exists_lt_of_tendsto_at_top hu _ _,
obtain ⟨n, hnN, hnM, hn_min⟩ : ∃ n, N ≤ n ∧ M < u n ∧ ∀ k, N ≤ k → k < n → u k ≤ M,
{ use nat.find ex,
rw ← and_assoc,
split,
{ simpa using nat.find_spec ex },
{ intros k hk hk',
simpa [hk] using nat.find_min ex hk' } },
use [n, hnN],
intros k hk,
by_cases H : k ≤ N,
{ have : u k ∈ A,
from finset.mem_image_of_mem _ (finset.mem_range.mpr $ nat.lt_succ_of_le H),
have : u k ≤ M,
from finset.le_max' A (u k) this,
exact lt_of_le_of_lt this hnM },
{ push_neg at H,
calc u k ≤ M : hn_min k (le_of_lt H) hk
... < u n : hnM },
obtain ⟨M, hMN, hMk, hM_min⟩ : ∃ M ≥ N, u k < u M ∧ ∀ m, m < M → N ≤ m → u m ≤ u k,
{ rcases nat.find_x ex with ⟨M, ⟨hMN, hMk⟩, hM_min⟩,
push_neg at hM_min,
exact ⟨M, hMN, hMk, hM_min⟩ },
refine ⟨M, hMN, λ l hl, lt_of_le_of_lt _ hMk⟩,
exact (le_total l N).elim (hku _) (hM_min l hl)
end

/--
Expand Down

0 comments on commit db9d4a3

Please sign in to comment.