Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - chore(data/finset/lattice): remove unneeded assumptions #4020

Closed
wants to merge 8 commits into from
Closed
Show file tree
Hide file tree
Changes from 7 commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Jump to
Jump to file
Failed to load files.
Diff view
Diff view
8 changes: 4 additions & 4 deletions src/combinatorics/composition.lean
Original file line number Diff line number Diff line change
Expand Up @@ -712,14 +712,14 @@ finset.mono_of_fin c.boundaries rfl

@[simp] lemma boundary_zero : c.boundary ⟨0, c.card_boundaries_pos⟩ = 0 :=
begin
rw [boundary, finset.mono_of_fin_zero rfl c.boundaries_nonempty c.card_boundaries_pos],
exact le_antisymm (finset.min'_le _ _ _ c.zero_mem) (fin.zero_le _),
rw [boundary, finset.mono_of_fin_zero rfl c.card_boundaries_pos],
exact le_antisymm (finset.min'_le _ _ c.zero_mem) (fin.zero_le _),
end

@[simp] lemma boundary_length : c.boundary ⟨c.length, c.length_lt_card_boundaries⟩ = fin.last n :=
begin
convert finset.mono_of_fin_last rfl c.boundaries_nonempty c.card_boundaries_pos,
exact le_antisymm (finset.le_max' _ _ _ c.last_mem) (fin.le_last _)
convert finset.mono_of_fin_last rfl c.card_boundaries_pos,
exact le_antisymm (finset.le_max' _ _ c.last_mem) (fin.le_last _)
end

/-- Size of the `i`-th block in a `composition_as_set`, seen as a function on `fin c.length`. -/
Expand Down
33 changes: 19 additions & 14 deletions src/data/finset/lattice.lean
Original file line number Diff line number Diff line change
Expand Up @@ -275,34 +275,37 @@ variables (s : finset α) (H : s.nonempty)

theorem min'_mem : s.min' H ∈ s := mem_of_min $ by simp [min']

theorem min'_le (x) (H2 : x ∈ s) : s.min' H ≤ x := min_le_of_mem H2 $ option.get_mem _
theorem min'_le (x) (H2 : x ∈ s) : s.min' ⟨x, H2⟩ ≤ x := min_le_of_mem H2 $ option.get_mem _

theorem le_min' (x) (H2 : ∀ y ∈ s, x ≤ y) : x ≤ s.min' H := H2 _ $ min'_mem _ _

/-- `{a}.min'` is `a`. -/
@[simp] lemma min'_singleton (a : α) {h} : ({a} : finset α).min' h = a :=
/-- `{a}.min' _` is `a`. -/
@[simp] lemma min'_singleton (a : α) :
({a} : finset α).min' (singleton_nonempty _) = a :=
by simp [min']

theorem max'_mem : s.max' H ∈ s := mem_of_max $ by simp [max']

theorem le_max' (x) (H2 : x ∈ s) : x ≤ s.max' H := le_max_of_mem H2 $ option.get_mem _
theorem le_max' (x) (H2 : x ∈ s) : x ≤ s.max' ⟨x, H2⟩ := le_max_of_mem H2 $ option.get_mem _

theorem max'_le (x) (H2 : ∀ y ∈ s, y ≤ x) : s.max' H ≤ x := H2 _ $ max'_mem _ _

/-- `{a}.max'` is `a`. -/
@[simp] lemma max'_singleton (a : α) {h} : ({a} : finset α).max' h = a :=
/-- `{a}.max' _` is `a`. -/
@[simp] lemma max'_singleton (a : α) :
({a} : finset α).max' (singleton_nonempty _) = a :=
by simp [max']

theorem min'_lt_max' {i j} (H1 : i ∈ s) (H2 : j ∈ s) (H3 : i ≠ j) : s.min' H < s.max' H :=
theorem min'_lt_max' {i j} (H1 : i ∈ s) (H2 : j ∈ s) (H3 : i ≠ j) :
sgouezel marked this conversation as resolved.
Show resolved Hide resolved
s.min' ⟨i, H1⟩ < s.max' ⟨i, H1⟩ :=
begin
rcases lt_trichotomy i j with H4 | H4 | H4,
{ have H5 := min'_le s H i H1,
have H6 := le_max' s H j H2,
{ have H5 := min'_le s i H1,
have H6 := le_max' s j H2,
apply lt_of_le_of_lt H5,
apply lt_of_lt_of_le H4 H6 },
{ cc },
{ have H5 := min'_le s H j H2,
have H6 := le_max' s H i H1,
{ have H5 := min'_le s j H2,
have H6 := le_max' s i H1,
apply lt_of_le_of_lt H5,
apply lt_of_lt_of_le H4 H6 }
end
Expand All @@ -311,15 +314,17 @@ end
If there's more than 1 element, the min' is less than the max'. An alternate version of
`min'_lt_max'` which is sometimes more convenient.
-/
lemma min'_lt_max'_of_card (h₂ : 1 < card s) : s.min' H < s.max' H :=
lemma min'_lt_max'_of_card (h₂ : 1 < card s) :
s.min' (finset.card_pos.mp $ lt_trans zero_lt_one h₂) <
s.max' (finset.card_pos.mp $ lt_trans zero_lt_one h₂) :=
begin
apply lt_of_not_ge,
intro a,
apply not_le_of_lt h₂ (le_of_eq _),
rw card_eq_one,
use max' s H,
use (max' s (finset.card_pos.mp $ lt_trans zero_lt_one h₂)),
rw eq_singleton_iff_unique_mem,
refine ⟨max'_mem _ _, λ t Ht, le_antisymm (le_max' s H t Ht) (le_trans a (min'_le s H t Ht))⟩,
refine ⟨max'_mem _ _, λ t Ht, le_antisymm (le_max' s t Ht) (le_trans a (min'_le s t Ht))⟩,
end

end max_min
Expand Down
40 changes: 21 additions & 19 deletions src/data/finset/sort.lean
Original file line number Diff line number Diff line change
Expand Up @@ -52,28 +52,31 @@ theorem sort_sorted_lt (s : finset α) :
list.sorted (<) (sort (≤) s) :=
(sort_sorted _ _).imp₂ (@lt_of_le_of_ne _ _) (sort_nodup _ _)

lemma sorted_zero_eq_min' (s : finset α) (h : 0 < (s.sort (≤)).length) (H : s.nonempty) :
(s.sort (≤)).nth_le 0 h = s.min' H :=
lemma sorted_zero_eq_min' (s : finset α) (h : 0 < s.card) :
(s.sort (≤)).nth_le 0 (by rwa length_sort) = s.min' (by rwa card_pos at h) :=
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

This might just be my inexperience, but are lemmas that use tactics in their statements easy to use?

Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

the presence of tactics is not the problem so much as the choice of a particular proof to go in that slot. Most tactics will ignore the hypothesis argument when matching up to defeq but occasionally it helps to have a free variable for the hypothesis. Having free variables on both sides of a rewrite can be a problem, though, since it doesn't actually pin down what the hypothesis in the RHS is, and so you get an extra subgoal you didn't want. So it's sometimes helpful to have both versions, or a version where the LHS has a variable and the RHS is concrete (if it is to be used for left to right rewrites or simps).

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

@digama0 So, what do you suggest? Should we revert, or continue?

Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Since the tactic is used to prove a Prop, it will match any Prop that is given to the lemma if you rewrite from left to right or from right to left, by proof irrelevance, and it will provide the required prop on the output of the rewrite. So, to me, this is an improvement over the previous situation, where the user had to provide more data. In all uses of this lemma, the outcome of the change is that we can remove one of the arguments to the call.

Mario, you say that occasionally it helps to have a free variable for the hypothesis. Do you have an example of this behavior?

Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I agree that the original form with two variables is not good for anyone, except in the case where you are matching an equality in the goal. But the new version of the lemma can never prove h by unification, while the original could, which means you are always going to have to prove the hypothesis h as a subgoal when you apply the new lemma.

Is there a definite common rewrite direction? If so, (say l-t-r) you can put the free hypothesis on the LHS (as is), and prove the RHS hypothesis using two rewrites on the LHS hypothesis. That way if you rewrite the hypothesis will be proven by unification, and you don't get any subgoal for the RHS because it's been proven in the theorem. 0 subgoals is better than 1, no?

If you don't know how the lemma will be used, I suggest providing both the original over-general theorem (this one getting a prime) and the present constrained theorem, and then users can use an appropriate version for the application. There are a few list lemmas like this.

Copy link
Collaborator

@agjftucker agjftucker Sep 3, 2020

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I think only a couple of these are used, and then only left-to-right. But if you are already considering two versions of each, then maybe one could work left-to-right and one right-to-left? Something like this:

lemma sorted_zero_eq_inf' (s : finset α) (h : 0 < (s.sort (≤)).length) : (s.sort (≤)).nth_le 0 h = s.1.inf' (by rwa [length_sort, card_pos] at h) := begin ... end

lemma inf'_eq_sorted_zero (s : finset α) (H : s.nonempty) : s.1.inf' H = (s.sort (≤)).nth_le 0 (by rwa [length_sort, card_pos]) := by rw sorted_zero_eq_inf'

lemma sorted_last_eq_sup' (s : finset α) (h : (s.sort (≤)).length - 1 < (s.sort (≤)).length) : (s.sort (≤)).nth_le ((s.sort (≤)).length - 1) h = s.1.sup' (by simpa [card_pos] using lt_of_le_of_lt (nat.zero_le _) h) := begin ... end

lemma sup'_eq_sorted_last (s : finset α) (H : s.nonempty) : s.1.sup' H = (s.sort (≤)).nth_le ((s.sort (≤)).length - 1) (by simpa using sub_lt (card_pos.mpr H) zero_lt_one) := by rw sorted_last_eq_sup'

lemma mono_of_fin_zero_eq_inf' {s : finset α} {k : ℕ} (h : s.card = k) (hz : 0 < k) : mono_of_fin s h ⟨0, hz⟩ = s.1.inf' (card_pos.1 (h.symm ▸ hz)) := begin ... end

lemma inf'_eq_mono_of_fin_zero {s : finset α} (hs : s.nonempty) : s.1.inf' hs = mono_of_fin s rfl ⟨0, card_pos.2 hs⟩ := by rw mono_of_fin_zero_eq_inf'

lemma mono_of_fin_last_eq_sup' {s : finset α} {k : ℕ} (h : s.card = k) (hz : 0 < k) : mono_of_fin s h ⟨k-1, sub_lt hz zero_lt_one⟩ = s.1.sup' (card_pos.1 (h.symm ▸ hz)) := begin ... end

lemma sup'_eq_mono_of_fin_last {s : finset α} (hs : s.nonempty) : s.1.sup' hs = mono_of_fin s rfl ⟨s.card - 1, sub_lt (card_pos.2 hs) zero_lt_one⟩ := by rw mono_of_fin_last_eq_sup' rfl (card_pos.2 hs)

lemma mono_of_fin_singleton_eq (a : α) (i : fin 1) {h} : mono_of_fin {a} h i = a := begin ... end

lemma eq_mono_of_fin_singleton (a : α) (i : fin 1) : a = mono_of_fin {a} (card_singleton a) i := by rw mono_of_fin_singleton_eq

begin
let l := s.sort (≤),
have H : s.nonempty, by rwa card_pos at h,
have h' : 0 < (s.sort (≤)).length, by rwa length_sort,
apply le_antisymm,
{ have : s.min' H ∈ l := (finset.mem_sort (≤)).mpr (s.min'_mem H),
obtain ⟨i, i_lt, hi⟩ : ∃ i (hi : i < l.length), l.nth_le i hi = s.min' H :=
list.mem_iff_nth_le.1 this,
rw ← hi,
exact list.nth_le_of_sorted_of_le (s.sort_sorted (≤)) (nat.zero_le i) },
{ have : l.nth_le 0 h ∈ s := (finset.mem_sort (≤)).1 (list.nth_le_mem l 0 h),
exact s.min'_le H _ this }
{ have : l.nth_le 0 h' ∈ s := (finset.mem_sort (≤)).1 (list.nth_le_mem l 0 h'),
exact s.min'_le _ this }
end

lemma sorted_last_eq_max' (s : finset α) (h : (s.sort (≤)).length - 1 < (s.sort (≤)).length)
(H : s.nonempty) : (s.sort (≤)).nth_le ((s.sort (≤)).length - 1) h = s.max' H :=
lemma sorted_last_eq_max' (s : finset α) (H : s.nonempty) :
(s.sort (≤)).nth_le ((s.sort (≤)).length - 1) (by simpa using sub_lt (card_pos.mpr H) zero_lt_one)
= s.max' H :=
begin
let l := s.sort (≤),
apply le_antisymm,
{ have : l.nth_le ((s.sort (≤)).length - 1) h ∈ s :=
(finset.mem_sort (≤)).1 (list.nth_le_mem l _ h),
exact s.le_max' H _ this },
{ have : l.nth_le ((s.sort (≤)).length - 1) _ ∈ s :=
(finset.mem_sort (≤)).1 (list.nth_le_mem l _ _),
exact s.le_max' _ this },
{ have : s.max' H ∈ l := (finset.mem_sort (≤)).mpr (s.max'_mem H),
obtain ⟨i, i_lt, hi⟩ : ∃ i (hi : i < l.length), l.nth_le i hi = s.max' H :=
list.mem_iff_nth_le.1 this,
Expand All @@ -90,7 +93,6 @@ def mono_of_fin (s : finset α) {k : ℕ} (h : s.card = k) (i : fin k) : α :=
have A : (i : ℕ) < (s.sort (≤)).length, by simpa [h] using i.2,
(s.sort (≤)).nth_le i A


lemma mono_of_fin_strict_mono (s : finset α) {k : ℕ} (h : s.card = k) :
strict_mono (s.mono_of_fin h) :=
begin
Expand Down Expand Up @@ -120,28 +122,28 @@ lemma mono_of_fin_injective (s : finset α) {k : ℕ} (h : s.card = k) :
set.injective_iff_inj_on_univ.mpr (s.mono_of_fin_bij_on h).inj_on

/-- The bijection `mono_of_fin s h` sends `0` to the minimum of `s`. -/
lemma mono_of_fin_zero {s : finset α} {k : ℕ} (h : s.card = k) (hs : s.nonempty) (hz : 0 < k) :
mono_of_fin s h ⟨0, hz⟩ = s.min' hs :=
lemma mono_of_fin_zero {s : finset α} {k : ℕ} (h : s.card = k) (hz : 0 < k) :
mono_of_fin s h ⟨0, hz⟩ = s.min' (card_pos.mp (h.symm ▸ hz)) :=
begin
apply le_antisymm,
{ have : min' s hs ∈ s := min'_mem s hs,
{ have : min' s _ ∈ s := min'_mem s _,
rcases (mono_of_fin_bij_on s h).surj_on this with ⟨a, _, ha⟩,
rw ← ha,
apply (mono_of_fin_strict_mono s h).monotone,
exact zero_le a.val },
{ have : mono_of_fin s h ⟨0, hz⟩ ∈ s := (mono_of_fin_bij_on s h).maps_to (set.mem_univ _),
exact min'_le s hs _ this }
exact min'_le s _ this }
end

/-- The bijection `mono_of_fin s h` sends `k-1` to the maximum of `s`. -/
lemma mono_of_fin_last {s : finset α} {k : ℕ} (h : s.card = k) (hs : s.nonempty) (hz : 0 < k) :
mono_of_fin s h ⟨k-1, buffer.lt_aux_2 hz⟩ = s.max' hs :=
lemma mono_of_fin_last {s : finset α} {k : ℕ} (h : s.card = k) (hz : 0 < k) :
mono_of_fin s h ⟨k-1, buffer.lt_aux_2 hz⟩ = s.max' (card_pos.mp (h.symm ▸ hz)) :=
begin
have h'' : k - 1 < k := buffer.lt_aux_2 hz,
apply le_antisymm,
{ have : mono_of_fin s h ⟨k-1, h''⟩ ∈ s := (mono_of_fin_bij_on s h).maps_to (set.mem_univ _),
exact le_max' s hs _ this },
{ have : max' s hs ∈ s := max'_mem s hs,
exact le_max' s _ this },
{ have : max' s _ ∈ s := max'_mem s _,
rcases (mono_of_fin_bij_on s h).surj_on this with ⟨a, _, ha⟩,
rw ← ha,
apply (mono_of_fin_strict_mono s h).monotone,
Expand All @@ -152,7 +154,7 @@ end
@[simp] lemma mono_of_fin_singleton (a : α) (i : fin 1) {h} :
mono_of_fin {a} h i = a :=
by rw [subsingleton.elim i ⟨0, zero_lt_one⟩,
mono_of_fin_zero h (singleton_nonempty a) zero_lt_one, min'_singleton]
mono_of_fin_zero h zero_lt_one, min'_singleton]

/-- The range of `mono_of_fin`. -/
@[simp] lemma range_mono_of_fin {s : finset α} {k : ℕ} (h : s.card = k) :
Expand Down
2 changes: 1 addition & 1 deletion src/order/filter/at_top_bot.lean
Original file line number Diff line number Diff line change
Expand Up @@ -291,7 +291,7 @@ begin
{ 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 Ane (u k) this,
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
Expand Down