]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/RELATIONAL/NLE/props.ma
some properties of NLE
[helm.git] / helm / software / matita / contribs / RELATIONAL / NLE / props.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/RELATIONAL/NLE/props".
16
17 include "NLE/fwd.ma".
18
19 theorem nle_refl: \forall x. x <= x.
20  intros 1; elim x; clear x; intros; auto.
21 qed.
22
23 theorem nle_trans_succ: \forall x,y. x <= y \to x <= succ y.
24  intros. elim H; clear H x y; intros; auto.
25 qed.
26
27 theorem nle_lt_or_eq: \forall y,x. x <= y \to x < y \lor x = y.
28  intros 1. elim y; clear y; intros;
29  [ lapply linear nle_gen_zero_2 to H. auto
30  | lapply linear nle_gen_succ_2 to H1. decompose;
31    [ rewrite > H1. clear H1. auto
32    | lapply linear H to H3 as H0. decompose;
33      [ rewrite > H1. clear H1 x. auto
34      | rewrite < H. clear H n. auto
35      ]
36    ]
37  ].
38 qed.