Skip to content
Snippets Groups Projects
Commit bb76b8bd authored by Pierre Roux's avatar Pierre Roux Committed by Björn Brandenburg
Browse files

Prepare subdnD for MathComp PR

parent b782b2d7
No related branches found
No related tags found
1 merge request!212various cleanups and simplifications in util
......@@ -6,16 +6,13 @@ Require Export prosa.util.tactics.
(** Additional lemmas about natural numbers. *)
Section NatLemmas.
(** First, we show that, given [m1 >= m2] and [n1 >= n2], an
expression [(m1 + n1) - (m2 + n2)] can be transformed into
expression [(m1 - m2) + (n1 - n2)]. *)
Lemma subnD:
forall m1 m2 n1 n2,
m1 >= m2 ->
n1 >= n2 ->
(m1 + n1) - (m2 + n2) = (m1 - m2) + (n1 - n2).
Proof. by ins; lia. Qed.
(** First, we show that, given [m >= p] and [n >= q], an
expression [(m + n) - (p + q)] can be transformed into
expression [(m - p) + (n - q)]. *)
Lemma subnACA m n p q : p <= m -> q <= n ->
(m + n) - (p + q) = (m - p) + (n - q).
Proof. by move=> plem qlen; rewrite subnDA addnBAC// addnBA// subnAC. Qed.
(** Next, we show that [m + p <= n] implies that [m <= n - p]. Note
that this lemma is similar to ssreflect's lemma [leq_subRL];
however, the current lemma has no precondition [n <= p], since it
......
......@@ -259,7 +259,7 @@ Section SumsOverSequences.
Proof.
move=> LEQ.
induction r; first by rewrite !big_nil subn0.
rewrite !big_cons subnD.
rewrite !big_cons subnACA.
- apply/eqP; rewrite eqn_add2l; apply/eqP; apply IHl.
by intros; apply LEQ; rewrite in_cons; apply/orP; right.
- by apply LEQ; rewrite in_cons; apply/orP; left.
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment