]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/RELATIONAL/NLE/props.ma
;auto fixed
[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_zero: \forall q. zero <= q.
20  unfold NLE.
21  intros. apply ex_intro; auto. (**)
22 qed.
23
24 theorem nle_succ: \forall p,q. p <= q \to succ p <= succ q.
25  unfold NLE.
26  intros. decompose.
27  apply ex_intro; [|auto] (**)
28 qed.
29
30 theorem nle_refl: \forall x. x <= x.
31  intros 1; elim x; clear x; intros; auto new timeout=100.
32 qed.
33
34 theorem nle_trans_succ: \forall x,y. x <= y \to x <= succ y.
35  intros 1. elim x; clear x; intros; 
36  [ auto new timeout=100.
37  | lapply linear nle_gen_succ_1 to H1 as H0. decompose H0. subst.
38    auto new timeout=100.
39  ].
40 qed.
41
42 theorem nle_lt_or_eq: \forall y,x. x <= y \to x < y \lor x = y.
43  intros 1. elim y; clear y; intros;
44  [ lapply linear nle_gen_zero_2 to H. auto new timeout=100
45  | lapply linear nle_gen_succ_2 to H1. decompose;
46    [ subst. auto new timeout=100
47    | lapply linear H to H3 as H0. decompose;
48      subst; auto new timeout=100
49    ]
50  ].
51 qed.