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)         *)
(************************************************************************)
Theorems about gt in nat.
This file is DEPRECATED now, see module PeanoNat.Nat instead, which favor lt over gt.
gt is defined in Init/Peano.v as:
Definition gt (n m:nat) := m < n.
Require Import PeanoNat Le Lt Plus.
Local Open Scope nat_scope.

Order and successor

n:nat

S n > 0
Proof Nat.lt_0_succ _.
n:nat

S n > n
Proof Nat.lt_succ_diag_r _.
n, m:nat

n > m -> S n > S m
n, m:nat

n > m -> S n > S m
apply Nat.succ_lt_mono. Qed.
n, m:nat

S m > S n -> m > n
n, m:nat

S m > S n -> m > n
apply Nat.succ_lt_mono. Qed.
n, m:nat

S n > m -> n > m \/ m = n
n, m:nat

S n > m -> n > m \/ m = n
n, m:nat
H:S n > m

n > m \/ m = n
now apply Nat.lt_eq_cases, Nat.succ_le_mono. Qed.
n, m:nat

m > S n -> Init.Nat.pred m > n
n, m:nat

m > S n -> Init.Nat.pred m > n
apply Nat.lt_succ_lt_pred. Qed.

Irreflexivity

n:nat

~ n > n
Proof Nat.lt_irrefl _.

Asymmetry

n, m:nat

n > m -> ~ m > n
Proof Nat.lt_asymm _ _.

Relating strict and large orders

n, m:nat

n <= m -> ~ n > m
n, m:nat

n <= m -> ~ n > m
apply Nat.le_ngt. Qed.
n, m:nat

n > m -> ~ n <= m
n, m:nat

n > m -> ~ n <= m
apply Nat.lt_nge. Qed.
n, m:nat

S n <= m -> m > n
n, m:nat

S n <= m -> m > n
apply Nat.le_succ_l. Qed.
n, m:nat

S m > n -> n <= m
n, m:nat

S m > n -> n <= m
apply Nat.succ_le_mono. Qed.
n, m:nat

m > n -> S n <= m
n, m:nat

m > n -> S n <= m
apply Nat.le_succ_l. Qed.
n, m:nat

n <= m -> S m > n
n, m:nat

n <= m -> S m > n
apply Nat.succ_le_mono. Qed.

Transitivity

n, m, p:nat

m <= n -> m > p -> n > p
n, m, p:nat

m <= n -> m > p -> n > p
n, m, p:nat
H:m <= n
H0:m > p

n > p
now apply Nat.lt_le_trans with m. Qed.
n, m, p:nat

n > m -> p <= m -> n > p
n, m, p:nat

n > m -> p <= m -> n > p
n, m, p:nat
H:n > m
H0:p <= m

n > p
now apply Nat.le_lt_trans with m. Qed.
n, m, p:nat

n > m -> m > p -> n > p
n, m, p:nat

n > m -> m > p -> n > p
n, m, p:nat
H:n > m
H0:m > p

n > p
now apply Nat.lt_trans with m. Qed.
n, m, p:nat

S n > m -> m > p -> n > p
n, m, p:nat

S n > m -> m > p -> n > p
n, m, p:nat
H:S n > m
H0:m > p

n > p
n, m, p:nat
H:S n > m
H0:m > p

m <= n
now apply Nat.succ_le_mono. Qed.

Comparison to 0

n:nat

n > 0 \/ 0 = n
n:nat

n > 0 \/ 0 = n
destruct n; [now right | left; apply Nat.lt_0_succ]. Qed.

Simplification and compatibility

n, m, p:nat

p + n > p + m -> n > m
n, m, p:nat

p + n > p + m -> n > m
apply Nat.add_lt_mono_l. Qed.
n, m, p:nat

n > m -> p + n > p + m
n, m, p:nat

n > m -> p + n > p + m
apply Nat.add_lt_mono_l. Qed.

Hints

Hint Resolve gt_Sn_O gt_Sn_n gt_n_S : arith.
Hint Immediate gt_S_n gt_pred : arith.
Hint Resolve gt_irrefl gt_asym : arith.
Hint Resolve le_not_gt gt_not_le : arith.
Hint Immediate le_S_gt gt_S_le : arith.
Hint Resolve gt_le_S le_gt_S : arith.
Hint Resolve gt_trans_S le_gt_trans gt_le_trans: arith.
Hint Resolve plus_gt_compat_l: arith.

(* begin hide *)
Notation gt_O_eq := gt_0_eq (only parsing).
(* end hide *)