summaryrefslogtreecommitdiff
path: root/theories/Arith/Lt.v
blob: e1b3e4b8cc4e918c66b1c7cac04b9c731146443d (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
(************************************************************************)
(*  v      *   The Coq Proof Assistant  /  The Coq Development Team     *)
(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
(*   \VV/  **************************************************************)
(*    //   *      This file is distributed under the terms of the       *)
(*         *       GNU Lesser General Public License Version 2.1        *)
(************************************************************************)

(*i $Id: Lt.v,v 1.11.2.1 2004/07/16 19:31:00 herbelin Exp $ i*)

Require Import Le.
Open Local Scope nat_scope.

Implicit Types m n p : nat.

(** Irreflexivity *)

Theorem lt_irrefl : forall n, ~ n < n.
Proof le_Sn_n.
Hint Resolve lt_irrefl: arith v62.

(** Relationship between [le] and [lt] *) 

Theorem lt_le_S : forall n m, n < m -> S n <= m.
Proof.
auto with arith.
Qed.
Hint Immediate lt_le_S: arith v62.

Theorem lt_n_Sm_le : forall n m, n < S m -> n <= m.
Proof.
auto with arith.
Qed.
Hint Immediate lt_n_Sm_le: arith v62.

Theorem le_lt_n_Sm : forall n m, n <= m -> n < S m.
Proof.
auto with arith.
Qed.
Hint Immediate le_lt_n_Sm: arith v62.

Theorem le_not_lt : forall n m, n <= m -> ~ m < n.
Proof.
induction 1; auto with arith.
Qed.

Theorem lt_not_le : forall n m, n < m -> ~ m <= n.
Proof.
red in |- *; intros n m Lt Le; exact (le_not_lt m n Le Lt).
Qed.
Hint Immediate le_not_lt lt_not_le: arith v62.

(** Asymmetry *)

Theorem lt_asym : forall n m, n < m -> ~ m < n.
Proof.
induction 1; auto with arith.
Qed.

(** Order and successor *)

Theorem lt_n_Sn : forall n, n < S n.
Proof.
auto with arith.
Qed.
Hint Resolve lt_n_Sn: arith v62.

Theorem lt_S : forall n m, n < m -> n < S m.
Proof.
auto with arith.
Qed.
Hint Resolve lt_S: arith v62.

Theorem lt_n_S : forall n m, n < m -> S n < S m.
Proof.
auto with arith.
Qed.
Hint Resolve lt_n_S: arith v62.

Theorem lt_S_n : forall n m, S n < S m -> n < m.
Proof.
auto with arith.
Qed.
Hint Immediate lt_S_n: arith v62.

Theorem lt_O_Sn : forall n, 0 < S n.
Proof.
auto with arith.
Qed.
Hint Resolve lt_O_Sn: arith v62.

Theorem lt_n_O : forall n, ~ n < 0.
Proof le_Sn_O.
Hint Resolve lt_n_O: arith v62.

(** Predecessor *)

Lemma S_pred : forall n m, m < n -> n = S (pred n).
Proof.
induction 1; auto with arith.
Qed.

Lemma lt_pred : forall n m, S n < m -> n < pred m.
Proof.
induction 1; simpl in |- *; auto with arith.
Qed.
Hint Immediate lt_pred: arith v62.

Lemma lt_pred_n_n : forall n, 0 < n -> pred n < n.
destruct 1; simpl in |- *; auto with arith.
Qed.
Hint Resolve lt_pred_n_n: arith v62.

(** Transitivity properties *)

Theorem lt_trans : forall n m p, n < m -> m < p -> n < p.
Proof.
induction 2; auto with arith.
Qed.

Theorem lt_le_trans : forall n m p, n < m -> m <= p -> n < p.
Proof.
induction 2; auto with arith.
Qed.

Theorem le_lt_trans : forall n m p, n <= m -> m < p -> n < p.
Proof.
induction 2; auto with arith.
Qed.

Hint Resolve lt_trans lt_le_trans le_lt_trans: arith v62.

(** Large = strict or equal *)

Theorem le_lt_or_eq : forall n m, n <= m -> n < m \/ n = m.
Proof.
induction 1; auto with arith.
Qed.

Theorem lt_le_weak : forall n m, n < m -> n <= m.
Proof.
auto with arith.
Qed.
Hint Immediate lt_le_weak: arith v62.

(** Dichotomy *)

Theorem le_or_lt : forall n m, n <= m \/ m < n.
Proof.
intros n m; pattern n, m in |- *; apply nat_double_ind; auto with arith.
induction 1; auto with arith.
Qed.

Theorem nat_total_order : forall n m, n <> m -> n < m \/ m < n.
Proof.
intros m n diff.
elim (le_or_lt n m); [ intro H'0 | auto with arith ].
elim (le_lt_or_eq n m); auto with arith.
intro H'; elim diff; auto with arith.
Qed.

(** Comparison to 0 *)

Theorem neq_O_lt : forall n, 0 <> n -> 0 < n.
Proof.
induction n; auto with arith.
intros; absurd (0 = 0); trivial with arith.
Qed.
Hint Immediate neq_O_lt: arith v62.

Theorem lt_O_neq : forall n, 0 < n -> 0 <> n.
Proof.
induction 1; auto with arith.
Qed.
Hint Immediate lt_O_neq: arith v62.