aboutsummaryrefslogtreecommitdiffhomepage
path: root/theories/Numbers/Integer/Axioms/ZOrder.v
blob: 3bf4d61f5a83c4525b1dbf857af18f311ef406ed (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
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
337
338
339
340
341
342
343
344
345
346
347
348
349
350
351
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
371
372
373
374
375
376
377
378
379
380
381
382
383
384
385
386
387
388
389
390
391
392
393
394
395
396
397
398
399
400
401
402
403
404
405
406
407
408
409
410
411
412
413
414
415
416
417
418
419
420
421
422
423
424
425
426
427
428
429
430
431
432
433
434
435
436
437
438
439
440
441
442
443
444
445
446
447
448
449
450
451
452
453
Require Import NumPrelude.
Require Import ZDomain.
Require Import ZAxioms.

Module Type OrderSignature.
Declare Module Export IntModule : IntSignature.
Open Local Scope ZScope.

Parameter Inline lt : Z -> Z -> bool.
Parameter Inline le : Z -> Z -> bool.
Add Morphism lt with signature E ==> E ==> eq_bool as lt_wd.
Add Morphism le with signature E ==> E ==> eq_bool as le_wd.

Notation "n <  m" := (lt n m) : ZScope.
Notation "n <= m" := (le n m) : ZScope.

Axiom le_lt : forall n m, n <= m <-> n < m \/ n == m.
Axiom lt_irr : forall n, ~ (n < n).
Axiom lt_S : forall n m, n < (S m) <-> n <= m.

End OrderSignature.


Module OrderProperties (Export OrderModule : OrderSignature).
Module Export IntPropertiesModule := IntProperties IntModule.
Open Local Scope ZScope.

Ltac le_intro1 := rewrite le_lt; left.
Ltac le_intro2 := rewrite le_lt; right.
Ltac le_elim H := rewrite le_lt in H; destruct H as [H | H].

Theorem le_refl : forall n, n <= n.
Proof.
intro n; now le_intro2.
Qed.

Theorem lt_n_Sn : forall n, n < S n.
Proof.
intro n. rewrite lt_S. now le_intro2.
Qed.

Theorem le_n_Sn : forall n, n <= S n.
Proof.
intro; le_intro1; apply lt_n_Sn.
Qed.

Theorem lt_Pn_n : forall n, P n < n.
Proof.
intro n; rewrite_S_P n at 2; apply lt_n_Sn.
Qed.

Theorem le_Pn_n : forall n, P n <= n.
Proof.
intro; le_intro1; apply lt_Pn_n.
Qed.

Theorem lt_n_Sm : forall n m, n < m -> n < S m.
Proof.
intros. rewrite lt_S. now le_intro1.
Qed.

Theorem le_n_Sm : forall n m, n <= m -> n <= S m.
Proof.
intros n m H; rewrite <- lt_S in H; now le_intro1.
Qed.

Theorem lt_n_m_P : forall n m, n < m <-> n <= P m.
Proof.
intros n m; rewrite_S_P m; rewrite P_S; apply lt_S.
Qed.

Theorem not_le_n_Pn : forall n, ~ n <= P n.
Proof.
intros n H; le_elim H.
apply lt_n_Sm in H; rewrite S_P in H; false_hyp H lt_irr.
pose proof (lt_Pn_n n) as H1; rewrite <- H in H1; false_hyp H1 lt_irr.
Qed.

Theorem le_S : forall n m, n <= S m <-> n <= m \/ n == S m.
Proof.
intros n m; rewrite le_lt. now rewrite lt_S.
Qed.

Theorem lt_P : forall n m, (P n) < m <-> n <= m.
Proof.
intro n; induct_n m (P n).
split; intro H. false_hyp H lt_irr. false_hyp H not_le_n_Pn.
intros m IH; split; intro H.
apply -> lt_S in H; le_elim H.
apply -> IH in H; now apply le_n_Sm.
rewrite <- H; rewrite S_P; now le_intro2.
apply -> le_S in H; destruct H as [H | H].
apply <- IH in H. now apply lt_n_Sm. rewrite H; rewrite P_S; apply lt_n_Sn.
intros m IH; split; intro H.
pose proof H as H1. apply lt_n_Sm in H; rewrite S_P in H.
apply -> IH in H; le_elim H. now apply -> lt_n_m_P.
rewrite H in H1; false_hyp H1 lt_irr.
pose proof H as H1. apply le_n_Sm in H. rewrite S_P in H.
apply <- IH in H. apply -> lt_n_m_P in H. le_elim H.
assumption. apply P_inj in H; rewrite H in H1; false_hyp H1 not_le_n_Pn.
Qed.

Theorem lt_Pn_m : forall n m, n < m -> P n < m.
Proof.
intros; rewrite lt_P; now le_intro1.
Qed.

Theorem le_Pn_m : forall n m, n <= m -> P n <= m.
Proof.
intros n m H; rewrite <- lt_P in H; now le_intro1.
Qed.

Theorem lt_n_m_S : forall n m, n < m <-> S n <= m.
Proof.
intros n m; rewrite_P_S n; rewrite S_P; apply lt_P.
Qed.

Theorem lt_Sn_m : forall n m, S n < m -> n < m.
Proof.
intros n m H; rewrite_P_S n; now apply lt_Pn_m.
Qed.

Theorem le_Sn_m : forall n m, S n <= m -> n <= m.
Proof.
intros n m H; rewrite <- lt_n_m_S in H; now le_intro1.
Qed.

Theorem lt_n_Pm : forall n m, n < P m -> n < m.
Proof.
intros n m H; rewrite_S_P m; now apply lt_n_Sm.
Qed.

Theorem le_n_Pm : forall n m, n <= P m -> n <= m.
Proof.
intros n m H; rewrite <- lt_n_m_P in H; now le_intro1.
Qed.

Theorem lt_respects_S : forall n m, n < m <-> S n < S m.
Proof.
intros n m. rewrite lt_n_m_S. symmetry. apply lt_S.
Qed.

Theorem le_respects_S : forall n m, n <= m <-> S n <= S m.
Proof.
intros n m. do 2 rewrite le_lt.
firstorder using lt_respects_S S_wd S_inj.
Qed.

Theorem lt_respects_P : forall n m, n < m <-> P n < P m.
Proof.
intros n m. rewrite lt_n_m_P. symmetry; apply lt_P.
Qed.

Theorem le_respects_P : forall n m, n <= m <-> P n <= P m.
Proof.
intros n m. do 2 rewrite le_lt.
firstorder using lt_respects_P P_wd P_inj.
Qed.

Theorem lt_S_P : forall n m, S n < m <-> n < P m.
Proof.
intros n m; rewrite_P_S n at 2; apply lt_respects_P.
Qed.

Theorem le_S_P : forall n m, S n <= m <-> n <= P m.
Proof.
intros n m; rewrite_P_S n at 2; apply le_respects_P.
Qed.

Theorem lt_P_S : forall n m, P n < m <-> n < S m.
Proof.
intros n m; rewrite_S_P n at 2; apply lt_respects_S.
Qed.

Theorem le_P_S : forall n m, P n <= m <-> n <= S m.
Proof.
intros n m; rewrite_S_P n at 2; apply le_respects_S.
Qed.

Theorem lt_neq : forall n m, n < m -> n # m.
Proof.
intros n m H1 H2; rewrite H2 in H1; false_hyp H1 lt_irr.
Qed.

Theorem lt_trans : forall n m p, n < m -> m < p -> n < p.
Proof.
intros n m; induct_n n m.
intros p H _; false_hyp H lt_irr.
intros n IH p H1 H2. apply lt_Sn_m in H1. pose proof (IH p H1 H2) as H3.
rewrite lt_n_m_S in H3; le_elim H3.
assumption. rewrite <- H3 in H2. rewrite lt_S in H2; le_elim H2.
elimtype False; apply lt_irr with (n := n); now apply IH.
rewrite H2 in H1; false_hyp H1 lt_irr.
intros n IH p H1 H2. apply lt_Pn_m. rewrite lt_P in H1; le_elim H1.
now apply IH. now rewrite H1.
Qed.

Theorem le_trans : forall n m p, n <= m -> m <= p -> n <= p.
Proof.
intros n m p H1 H2; le_elim H1.
le_elim H2. le_intro1; now apply lt_trans with (m := m).
le_intro1; now rewrite <- H2. now rewrite H1.
Qed.

Theorem le_lt_trans : forall n m p, n <= m -> m < p -> n < p.
Proof.
intros n m p H1 H2; le_elim H1.
now apply lt_trans with (m := m). now rewrite H1.
Qed.

Theorem lt_le_trans : forall n m p, n < m -> m <= p -> n < p.
Proof.
intros n m p H1 H2; le_elim H2.
now apply lt_trans with (m := m). now rewrite <- H2.
Qed.

Theorem lt_asymm : forall n m, n < m -> ~ m < n.
Proof.
intros n m H1 H2; apply lt_irr with (n := n); now apply lt_trans with (m := m).
Qed.

Theorem le_antisym : forall n m, n <= m -> m <= n -> n == m.
Proof.
intros n m H1 H2; le_elim H1; le_elim H2.
elimtype False; apply lt_irr with (n := n); now apply lt_trans with (m := m).
now symmetry. assumption. assumption.
Qed.

Theorem not_lt_Sn_n : forall n, ~ S n < n.
Proof.
intros n H; apply (lt_asymm n (S n)). apply lt_n_Sn. assumption.
Qed.

Theorem not_le_Sn_n : forall n, ~ S n <= n.
Proof.
intros n H; le_elim H. false_hyp H not_lt_Sn_n.
pose proof (lt_n_Sn n) as H1. rewrite H in H1; false_hyp H1 lt_irr.
Qed.

Theorem lt_gt : forall n m, n < m -> m < n -> False.
Proof.
intros n m H1 H2; apply lt_irr with (n := n); now apply lt_trans with (m := m).
Qed.

Theorem lt_total : forall n m,  n < m \/ n == m \/ m < n.
Proof.
intros n m; induct_n n m.
right; now left.
intros n IH; destruct IH as [H | [H | H]].
rewrite lt_n_m_S in H. rewrite le_lt in H; tauto.
right; right; rewrite H; apply lt_n_Sn.
right; right; now apply lt_n_Sm.
intros n IH; destruct IH as [H | [H | H]].
left; now apply lt_Pn_m.
left; rewrite H; apply lt_Pn_n.
rewrite lt_n_m_P in H. rewrite le_lt in H.
setoid_replace (m == P n) with (P n == m) in H using relation iff. tauto.
split; intro; now symmetry.
Qed.

Theorem le_gt : forall n m, n <= m <-> ~ m < n.
Proof.
intros n m. rewrite -> le_lt.
pose proof (lt_total n m). pose proof (lt_gt n m).
assert (n == m -> ~ m < n); [intro A; rewrite A; apply lt_irr |].
tauto.
Qed.

Theorem lt_ge : forall n m, n < m <-> ~ m <= n.
Proof.
intros n m. rewrite -> le_lt.
pose proof (lt_total m n). pose proof (lt_gt n m).
assert (n < m -> m # n); [intros A B; rewrite B in A; false_hyp A lt_irr |].
tauto.
Qed.

Theorem lt_discrete : forall n m, n < m -> m < S n -> False.
Proof.
intros n m H1 H2; apply -> lt_S in H2; apply -> lt_ge in H1; false_hyp H2 H1.
Qed.

(* Decidability of order can be proved either from totality or from the fact
that < and <= are boolean functions *)

(** A corollary of having an order is that Z is infinite in both
directions *)

Theorem neq_Sn_n : forall n, S n # n.
Proof.
intros n H. pose proof (lt_n_Sn n) as H1. rewrite H in H1. false_hyp H1 lt_irr.
Qed.

Theorem neq_Pn_n : forall n, P n # n.
Proof.
intros n H. apply S_wd in H. rewrite S_P in H. now apply neq_Sn_n with (n := n).
Qed.

Definition nth_S (n : nat) (m : Z) :=
  nat_rec (fun _ => Z) m (fun _ l => S l) n.
Definition nth_P (n : nat) (m : Z) :=
  nat_rec (fun _ => Z) m (fun _ l => P l) n.

Lemma lt_m_Skm :
  forall (n : nat) (m : Z), m < nth_S (Datatypes.S n) m.
Proof.
intros n m; induction n as [| n IH]; simpl in *.
apply lt_n_Sn. now apply lt_n_Sm.
Qed.

Lemma lt_Pkm_m :
  forall (n : nat) (m : Z), nth_P (Datatypes.S n) m < m.
Proof.
intros n m; induction n as [| n IH]; simpl in *.
apply lt_Pn_n. now apply lt_Pn_m.
Qed.

Theorem neq_m_Skm :
  forall (n : nat) (m : Z), nth_S (Datatypes.S n) m # m.
Proof.
intros n m H. pose proof (lt_m_Skm n m) as H1. rewrite H in H1.
false_hyp H1 lt_irr.
Qed.

Theorem neq_Pkm_m :
  forall (n : nat) (m : Z), nth_P (Datatypes.S n) m # m.
Proof.
intros n m H. pose proof (lt_Pkm_m n m) as H1. rewrite H in H1.
false_hyp H1 lt_irr.
Qed.

(** Stronger variant of induction with assumptions n >= 0 (n <= 0)
in the induction step *)

Section Induction.

Variable Q : Z -> Prop.
Hypothesis Q_wd : pred_wd E Q.

Add Morphism Q with signature E ==> iff as Q_morph.
Proof Q_wd.

Section Center.

Variable z : Z. (* Q z is the basis of induction *)

Section RightInduction.

Let Q' := fun n : Z => forall m, z <= m -> m < n -> Q m.

Add Morphism Q' with signature E ==> iff as Q'_pos_wd.
Proof.
intros x1 x2 H; unfold Q'; qmorphism x1 x2.
Qed.

Theorem right_induction :
  Q z -> (forall n, z <= n -> Q n -> Q (S n)) -> forall n, z <= n -> Q n.
Proof.
intros Qz QS k k_ge_z.
assert (H : forall n, Q' n). induct_n n z; unfold Q'.
intros m H1 H2. apply -> le_gt in H1; false_hyp H2 H1.
intros n IH m H2 H3.
rewrite lt_S in H3; le_elim H3. now apply IH.
le_elim H2. rewrite_S_P m.
apply QS. now apply -> lt_n_m_P. apply IH. now apply -> lt_n_m_P.
rewrite H3; apply lt_Pn_n. now rewrite <- H2.
intros n IH m H2 H3. apply IH. assumption. now apply lt_n_Pm.
pose proof (H (S k)) as H1; unfold Q' in H1. apply H1.
apply k_ge_z. apply lt_n_Sn.
Qed.

End RightInduction.

Section LeftInduction.

Let Q' := fun n : Z => forall m, m <= z -> n < m -> Q m.

Add Morphism Q' with signature E ==> iff as Q'_neg_wd.
Proof.
intros x1 x2 H; unfold Q'; qmorphism x1 x2.
Qed.

Theorem left_induction :
  Q z -> (forall n, n <= z -> Q n -> Q (P n)) -> forall n, n <= z -> Q n.
Proof.
intros Qz QP k k_le_z.
assert (H : forall n, Q' n). induct_n n z; unfold Q'.
intros m H1 H2. apply -> le_gt in H1; false_hyp H2 H1.
intros n IH m H2 H3. apply IH. assumption. now apply lt_Sn_m.
intros n IH m H2 H3.
rewrite lt_P in H3; le_elim H3. now apply IH.
le_elim H2. rewrite_P_S m.
apply QP. now apply -> lt_n_m_S. apply IH. now apply -> lt_n_m_S.
rewrite H3; apply lt_n_Sn. now rewrite H2.
pose proof (H (P k)) as H1; unfold Q' in H1. apply H1.
apply k_le_z. apply lt_Pn_n.
Qed.

End LeftInduction.

Theorem induction_ord_n :
  Q z ->
  (forall n, z <= n -> Q n -> Q (S n)) ->
  (forall n, n <= z -> Q n -> Q (P n)) ->
    forall n, Q n.
Proof.
intros Qz QS QP n.
destruct (lt_total n z) as [H | [H | H]].
now apply left_induction; [| | le_intro1].
now rewrite H.
now apply right_induction; [| | le_intro1].
Qed.

End Center.

Theorem induction_ord :
  Q 0 ->
  (forall n, 0 <= n -> Q n -> Q (S n)) ->
  (forall n, n <= 0 -> Q n -> Q (P n)) ->
    forall n, Q n.
Proof (induction_ord_n 0).

Theorem lt_ind : forall (n : Z),
  Q (S n) ->
  (forall m : Z, n < m -> Q m -> Q (S m)) ->
   forall m : Z, n < m -> Q m.
Proof.
intros n H1 H2 m H3.
apply right_induction with (S n). assumption.
intros; apply H2; try assumption. now apply <- lt_n_m_S.
now apply -> lt_n_m_S.
Qed.

Theorem le_ind : forall (n : Z),
  Q n ->
  (forall m : Z, n <= m -> Q m -> Q (S m)) ->
   forall m : Z, n <= m -> Q m.
Proof.
intros n H1 H2 m H3.
now apply right_induction with n.
Qed.

End Induction.

Ltac induct_ord n :=
  try intros until n;
  pattern n; apply induction_ord; clear n;
  [unfold NumPrelude.pred_wd;
  let n := fresh "n" in
  let m := fresh "m" in
  let H := fresh "H" in intros n m H; qmorphism n m | | |].

End OrderProperties.