]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_le.ma
- natural numbers with infinity for lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / ynat / ynat_le.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 include "ground_2/ynat/ynat_succ.ma".
16
17 (* NATURAL NUMBERS WITH INFINITY ********************************************)
18
19 (* order relation *)
20 inductive yle: relation ynat ≝
21 | yle_inj: ∀m,n. m ≤ n → yle m n
22 | yle_Y  : ∀m. yle m (∞)
23 .
24
25 interpretation "ynat 'less or equal to'" 'leq x y = (yle x y).
26
27 (* Basic inversion lemmas ***************************************************)
28
29 fact yle_inv_inj_aux: ∀x,y. x ≤ y → ∀m,n. x = yinj m → y = yinj n → m ≤ n.
30 #x #y * -x -y
31 [ #x #y #Hxy #m #n #Hx #Hy destruct //
32 | #x #m #n #_ #Hy destruct
33 ]
34 qed-.
35
36 lemma yle_inv_inj: ∀m,n. yinj m ≤ yinj n → m ≤ n.
37 /2 width=5 by yle_inv_inj_aux/ qed-.
38
39 fact yle_inv_O2_aux: ∀m:ynat. ∀x:ynat. m ≤ x → x = 0 → m = 0.
40 #m #x * -m -x
41 [ #m #n #Hmn #H destruct /3 width=1 by le_n_O_to_eq, eq_f/
42 | #m #H destruct
43
44 qed-.
45
46 lemma yle_inv_O2: ∀m:ynat. m ≤ 0 → m = 0.
47 /2 width =3 by yle_inv_O2_aux/ qed-.
48
49 fact yle_inv_Y1_aux: ∀x,n. x ≤ n → x = ∞ → n = ∞.
50 #x #n * -x -n //
51 #x #n #_ #H destruct
52 qed-.
53
54 lemma yle_inv_Y1: ∀n. ∞ ≤ n → n = ∞.
55 /2 width=3 by yle_inv_Y1_aux/ qed-.
56
57 (* Inversion lemmas on successor ********************************************)
58
59 fact yle_inv_succ1_aux: ∀x,y. x ≤ y → ∀m. x = ⫯m → ∃∃n. m ≤ n & y = ⫯n.
60 #x #y * -x -y
61 [ #x #y #Hxy #m #H elim (ysucc_inv_inj_sn … H) -H
62   #n #H1 #H2 destruct elim (le_inv_S1 … Hxy) -Hxy
63   #m #Hnm #H destruct
64   @(ex2_intro … m) /2 width=1 by yle_inj/ (**) (* explicit constructor *)
65 | #x #y #H destruct
66   @(ex2_intro … (∞)) /2 width=1 by yle_Y/
67 ]
68 qed-.
69
70 lemma yle_inv_succ1: ∀m,y.  ⫯m ≤ y → ∃∃n. m ≤ n & y = ⫯n.
71 /2 width=3 by yle_inv_succ1_aux/ qed-.
72
73 lemma yle_inv_succ: ∀m,n. ⫯m ≤ ⫯n → m ≤ n.
74 #m #n #H elim (yle_inv_succ1 … H) -H
75 #x #Hx #H destruct //
76 qed-.
77
78 (* Basic properties *********************************************************)
79
80 lemma yle_refl: reflexive … yle.
81 * /2 width=1 by le_n, yle_inj/
82 qed.
83
84 (* Properties on predecessor ************************************************)
85
86 lemma yle_pred_sn: ∀m,n. m ≤ n → ⫰m ≤ n.
87 #m #n * -m -n /3 width=3 by transitive_le, yle_inj/
88 qed.
89
90 lemma yle_refl_pred_sn: ∀x. ⫰x ≤ x.
91 /2 width=1 by yle_refl, yle_pred_sn/ qed.
92
93 (* Properties on successor **************************************************)
94
95 lemma yle_succ_dx: ∀m,n. m ≤ n → m ≤ ⫯n.
96 #m #n * -m -n /3 width=1 by le_S, yle_inj/
97 qed.
98
99 lemma yle_refl_S_dx: ∀x. x ≤ ⫯x.
100 /2 width=1 by yle_refl, yle_succ_dx/ qed.
101
102 (* Main properties **********************************************************)
103
104 theorem yle_trans: Transitive … yle.
105 #x #y * -x -y
106 [ #x #y #Hxy * //
107   #z #H lapply (yle_inv_inj … H) -H
108   /3 width=3 by transitive_le, yle_inj/ (**) (* full auto too slow *)
109 | #x #z #H lapply ( yle_inv_Y1 … H) //
110 ]
111 qed-.