summaryrefslogtreecommitdiff
path: root/theories/Arith/Peano_dec.v
blob: 340a7968915079d8e5f50f194e472e1ccea5c3d6 (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
(************************************************************************)
(*  v      *   The Coq Proof Assistant  /  The Coq Development Team     *)
(* <O___,, *   INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016     *)
(*   \VV/  **************************************************************)
(*    //   *      This file is distributed under the terms of the       *)
(*         *       GNU Lesser General Public License Version 2.1        *)
(************************************************************************)

Require Import Decidable PeanoNat.
Require Eqdep_dec.
Local Open Scope nat_scope.

Implicit Types m n x y : nat.

Theorem O_or_S n : {m : nat | S m = n} + {0 = n}.
Proof.
  induction n.
  - now right.
  - left; exists n; auto.
Defined.

Notation eq_nat_dec := Nat.eq_dec (compat "8.4").

Hint Resolve O_or_S eq_nat_dec: arith.

Theorem dec_eq_nat n m : decidable (n = m).
Proof.
  elim (Nat.eq_dec n m); [left|right]; trivial.
Defined.

Definition UIP_nat:= Eqdep_dec.UIP_dec Nat.eq_dec.

Import EqNotations.

Lemma le_unique: forall m n (le_mn1 le_mn2 : m <= n), le_mn1 = le_mn2.
Proof.
intros m n.
generalize (eq_refl (S n)).
generalize n at -1.
induction (S n) as [|n0 IHn0]; try discriminate.
clear n; intros n H; injection H; clear H; intro H.
rewrite <- H; intros le_mn1 le_mn2; clear n H.
pose (def_n2 := eq_refl n0); transitivity (eq_ind _ _ le_mn2 _ def_n2).
  2: reflexivity.
generalize def_n2; revert le_mn1 le_mn2.
generalize n0 at 1 4 5 7; intros n1 le_mn1.
destruct le_mn1; intros le_mn2; destruct le_mn2.
+ now intros def_n0; rewrite (UIP_nat _ _ def_n0 eq_refl).
+ intros def_n0; generalize le_mn2; rewrite <-def_n0; intros le_mn0.
  now destruct (Nat.nle_succ_diag_l _ le_mn0).
+ intros def_n0; generalize le_mn1; rewrite def_n0; intros le_mn0.
  now destruct (Nat.nle_succ_diag_l _ le_mn0).
+ intros def_n0; injection def_n0; intros ->.
  rewrite (UIP_nat _ _ def_n0 eq_refl); simpl.
  assert (H : le_mn1 = le_mn2).
    now apply IHn0.
now rewrite H.
Qed.

(** For compatibility *)
Require Import Le Lt.