Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) (* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) (* // * This file is distributed under the terms of the *) (* * GNU Lesser General Public License Version 2.1 *) (* * (see LICENSE file for the text of the license) *) (************************************************************************) Require Import NAxioms NSub GenericMinMax.
Module Type NMaxMinProp (Import N : NAxiomsMiniSig'). Include NSubProp N.
Zero
forall n : t, max 0 n == nforall n : t, max 0 n == nn:tmax 0 n == napply le_0_l. Qed.n:t0 <= nforall n : t, max n 0 == nforall n : t, max n 0 == nn:tmax n 0 == napply le_0_l. Qed.n:t0 <= nforall n : t, min 0 n == 0forall n : t, min 0 n == 0n:tmin 0 n == 0apply le_0_l. Qed.n:t0 <= nforall n : t, min n 0 == 0forall n : t, min n 0 == 0n:tmin n 0 == 0apply le_0_l. Qed.n:t0 <= n
The following results are concrete instances of max_monotone
and similar lemmas.
Succ
forall n m : t, S (max n m) == max (S n) (S m)forall n m : t, S (max n m) == max (S n) (S m)destruct (le_ge_cases n m); [rewrite 2 max_r | rewrite 2 max_l]; now rewrite <- ?succ_le_mono. Qed.n, m:tS (max n m) == max (S n) (S m)forall n m : t, S (min n m) == min (S n) (S m)forall n m : t, S (min n m) == min (S n) (S m)destruct (le_ge_cases n m); [rewrite 2 min_l | rewrite 2 min_r]; now rewrite <- ?succ_le_mono. Qed.n, m:tS (min n m) == min (S n) (S m)
Add
forall n m p : t, max (p + n) (p + m) == p + max n mforall n m p : t, max (p + n) (p + m) == p + max n mdestruct (le_ge_cases n m); [rewrite 2 max_r | rewrite 2 max_l]; now rewrite <- ?add_le_mono_l. Qed.n, m, p:tmax (p + n) (p + m) == p + max n mforall n m p : t, max (n + p) (m + p) == max n m + pforall n m p : t, max (n + p) (m + p) == max n m + pdestruct (le_ge_cases n m); [rewrite 2 max_r | rewrite 2 max_l]; now rewrite <- ?add_le_mono_r. Qed.n, m, p:tmax (n + p) (m + p) == max n m + pforall n m p : t, min (p + n) (p + m) == p + min n mforall n m p : t, min (p + n) (p + m) == p + min n mdestruct (le_ge_cases n m); [rewrite 2 min_l | rewrite 2 min_r]; now rewrite <- ?add_le_mono_l. Qed.n, m, p:tmin (p + n) (p + m) == p + min n mforall n m p : t, min (n + p) (m + p) == min n m + pforall n m p : t, min (n + p) (m + p) == min n m + pdestruct (le_ge_cases n m); [rewrite 2 min_l | rewrite 2 min_r]; now rewrite <- ?add_le_mono_r. Qed.n, m, p:tmin (n + p) (m + p) == min n m + p
Mul
forall n m p : t, max (p * n) (p * m) == p * max n mforall n m p : t, max (p * n) (p * m) == p * max n mdestruct (le_ge_cases n m); [rewrite 2 max_r | rewrite 2 max_l]; try order; now apply mul_le_mono_l. Qed.n, m, p:tmax (p * n) (p * m) == p * max n mforall n m p : t, max (n * p) (m * p) == max n m * pforall n m p : t, max (n * p) (m * p) == max n m * pdestruct (le_ge_cases n m); [rewrite 2 max_r | rewrite 2 max_l]; try order; now apply mul_le_mono_r. Qed.n, m, p:tmax (n * p) (m * p) == max n m * pforall n m p : t, min (p * n) (p * m) == p * min n mforall n m p : t, min (p * n) (p * m) == p * min n mdestruct (le_ge_cases n m); [rewrite 2 min_l | rewrite 2 min_r]; try order; now apply mul_le_mono_l. Qed.n, m, p:tmin (p * n) (p * m) == p * min n mforall n m p : t, min (n * p) (m * p) == min n m * pforall n m p : t, min (n * p) (m * p) == min n m * pdestruct (le_ge_cases n m); [rewrite 2 min_l | rewrite 2 min_r]; try order; now apply mul_le_mono_r. Qed.n, m, p:tmin (n * p) (m * p) == min n m * p
Sub
forall n m p : t, max (p - n) (p - m) == p - min n mforall n m p : t, max (p - n) (p - m) == p - min n mn, m, p:tmax (p - n) (p - m) == p - min n mn, m, p:tH:n <= mmax (p - n) (p - m) == p - min n mn, m, p:tH:m <= nmax (p - n) (p - m) == p - min n mn, m, p:tH:n <= mmax (p - n) (p - m) == p - nn, m, p:tH:m <= nmax (p - n) (p - m) == p - min n mn, m, p:tH:n <= mp - m <= p - nn, m, p:tH:m <= nmax (p - n) (p - m) == p - min n mn, m, p:tH:m <= nmax (p - n) (p - m) == p - min n mn, m, p:tH:m <= nmax (p - n) (p - m) == p - mnow apply sub_le_mono_l. Qed.n, m, p:tH:m <= np - n <= p - mforall n m p : t, max (n - p) (m - p) == max n m - pforall n m p : t, max (n - p) (m - p) == max n m - pdestruct (le_ge_cases n m); [rewrite 2 max_r | rewrite 2 max_l]; try order; now apply sub_le_mono_r. Qed.n, m, p:tmax (n - p) (m - p) == max n m - pforall n m p : t, min (p - n) (p - m) == p - max n mforall n m p : t, min (p - n) (p - m) == p - max n mn, m, p:tmin (p - n) (p - m) == p - max n mn, m, p:tH:n <= mmin (p - n) (p - m) == p - max n mn, m, p:tH:m <= nmin (p - n) (p - m) == p - max n mn, m, p:tH:n <= mmin (p - n) (p - m) == p - mn, m, p:tH:m <= nmin (p - n) (p - m) == p - max n mn, m, p:tH:n <= mp - m <= p - nn, m, p:tH:m <= nmin (p - n) (p - m) == p - max n mn, m, p:tH:m <= nmin (p - n) (p - m) == p - max n mn, m, p:tH:m <= nmin (p - n) (p - m) == p - nnow apply sub_le_mono_l. Qed.n, m, p:tH:m <= np - n <= p - mforall n m p : t, min (n - p) (m - p) == min n m - pforall n m p : t, min (n - p) (m - p) == min n m - pdestruct (le_ge_cases n m); [rewrite 2 min_l | rewrite 2 min_r]; try order; now apply sub_le_mono_r. Qed. End NMaxMinProp.n, m, p:tmin (n - p) (m - p) == min n m - p