Skip to content

Commit

Permalink
feat(algebra/big_operators/order): add fintype.sum_mono and fintype.s…
Browse files Browse the repository at this point in the history
…um_strict_mono (#6040)
  • Loading branch information
eric-wieser committed Feb 5, 2021
1 parent a101788 commit 80c7ac1
Showing 1 changed file with 13 additions and 0 deletions.
13 changes: 13 additions & 0 deletions src/algebra/big_operators/order.lean
Expand Up @@ -317,6 +317,19 @@ end canonically_ordered_comm_semiring

end finset

namespace fintype

variables [fintype α]

@[mono] lemma sum_mono [ordered_add_comm_monoid β] : monotone (λ f : α → β, ∑ x, f x) :=
λ f g hfg, finset.sum_le_sum $ λ x _, hfg x

lemma sum_strict_mono [ordered_cancel_add_comm_monoid β] : strict_mono (λ f : α → β, ∑ x, f x) :=
λ f g hfg, let ⟨hle, i, hlt⟩ := pi.lt_def.mp hfg in
finset.sum_lt_sum (λ i _, hle i) ⟨i, finset.mem_univ i, hlt⟩

end fintype

namespace with_top
open finset

Expand Down

0 comments on commit 80c7ac1

Please sign in to comment.