]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_le.ma
- improved arithmetics for natural numbers with infinity
[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_inj2_aux: ∀x,y. x ≤ y → ∀n. y = yinj n →
30                        ∃∃m. m ≤ n & x = yinj m.
31 #x #y * -x -y
32 [ #x #y #Hxy #n #Hy destruct /2 width=3 by ex2_intro/
33 | #x #n #Hy destruct
34 ]
35 qed-.
36
37 lemma yle_inv_inj2: ∀x,n. x ≤ yinj n → ∃∃m. m ≤ n & x = yinj m.
38 /2 width=3 by yle_inv_inj2_aux/ qed-.
39
40 lemma yle_inv_inj: ∀m,n. yinj m ≤ yinj n → m ≤ n.
41 #m #n #H elim (yle_inv_inj2 … H) -H
42 #x #Hxn #H destruct //
43 qed-.
44
45 fact yle_inv_O2_aux: ∀m:ynat. ∀x:ynat. m ≤ x → x = 0 → m = 0.
46 #m #x * -m -x
47 [ #m #n #Hmn #H destruct /3 width=1 by le_n_O_to_eq, eq_f/
48 | #m #H destruct
49
50 qed-.
51
52 lemma yle_inv_O2: ∀m:ynat. m ≤ 0 → m = 0.
53 /2 width =3 by yle_inv_O2_aux/ qed-.
54
55 fact yle_inv_Y1_aux: ∀x,n. x ≤ n → x = ∞ → n = ∞.
56 #x #n * -x -n //
57 #x #n #_ #H destruct
58 qed-.
59
60 lemma yle_inv_Y1: ∀n. ∞ ≤ n → n = ∞.
61 /2 width=3 by yle_inv_Y1_aux/ qed-.
62
63 (* Inversion lemmas on successor ********************************************)
64
65 fact yle_inv_succ1_aux: ∀x,y. x ≤ y → ∀m. x = ⫯m → ∃∃n. m ≤ n & y = ⫯n.
66 #x #y * -x -y
67 [ #x #y #Hxy #m #H elim (ysucc_inv_inj_sn … H) -H
68   #n #H1 #H2 destruct elim (le_inv_S1 … Hxy) -Hxy
69   #m #Hnm #H destruct
70   @(ex2_intro … m) /2 width=1 by yle_inj/ (**) (* explicit constructor *)
71 | #x #y #H destruct
72   @(ex2_intro … (∞)) /2 width=1 by yle_Y/ (**) (* explicit constructor *)
73 ]
74 qed-.
75
76 lemma yle_inv_succ1: ∀m,y.  ⫯m ≤ y → ∃∃n. m ≤ n & y = ⫯n.
77 /2 width=3 by yle_inv_succ1_aux/ qed-.
78
79 lemma yle_inv_succ: ∀m,n. ⫯m ≤ ⫯n → m ≤ n.
80 #m #n #H elim (yle_inv_succ1 … H) -H
81 #x #Hx #H destruct //
82 qed-.
83
84 (* Forward lemmas on successor **********************************************)
85
86 lemma yle_fwd_succ1: ∀m,n. ⫯m ≤ n → m ≤ ⫰n.
87 #m #x #H elim (yle_inv_succ1 … H) -H
88 #n #Hmn #H destruct //
89 qed-.
90
91 (* Basic properties *********************************************************)
92
93 lemma le_O1: ∀n:ynat. 0 ≤ n.
94 * /2 width=1 by yle_inj/
95 qed.
96
97 lemma yle_refl: reflexive … yle.
98 * /2 width=1 by le_n, yle_inj/
99 qed.
100
101 (* Properties on predecessor ************************************************)
102
103 lemma yle_pred_sn: ∀m,n. m ≤ n → ⫰m ≤ n.
104 #m #n * -m -n /3 width=3 by transitive_le, yle_inj/
105 qed.
106
107 lemma yle_refl_pred_sn: ∀x. ⫰x ≤ x.
108 /2 width=1 by yle_refl, yle_pred_sn/ qed.
109
110 (* Properties on successor **************************************************)
111
112 lemma yle_succ: ∀m,n. m ≤ n → ⫯m ≤ ⫯n.
113 #m #n * -m -n /3 width=1 by yle_inj, le_S_S/
114 qed.
115
116 lemma yle_succ_dx: ∀m,n. m ≤ n → m ≤ ⫯n.
117 #m #n * -m -n /3 width=1 by le_S, yle_inj/
118 qed.
119
120 lemma yle_refl_S_dx: ∀x. x ≤ ⫯x.
121 /2 width=1 by yle_refl, yle_succ_dx/ qed.
122
123 (* Main properties **********************************************************)
124
125 theorem yle_trans: Transitive … yle.
126 #x #y * -x -y
127 [ #x #y #Hxy * //
128   #z #H lapply (yle_inv_inj … H) -H
129   /3 width=3 by transitive_le, yle_inj/ (**) (* full auto too slow *)
130 | #x #z #H lapply (yle_inv_Y1 … H) //
131 ]
132 qed-.