]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/compare.ma
Towards chebyshev.
[helm.git] / matita / library / nat / compare.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||       A.Asperti, C.Sacerdoti Coen,                          *)
8 (*      ||A||       E.Tassi, S.Zacchiroli                                 *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU Lesser General Public License Version 2.1         *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/nat/compare".
16
17 include "datatypes/bool.ma".
18 include "datatypes/compare.ma".
19 include "nat/orders.ma".
20
21 let rec eqb n m \def 
22 match n with 
23   [ O \Rightarrow 
24      match m with 
25      [ O \Rightarrow true
26            | (S q) \Rightarrow false] 
27   | (S p) \Rightarrow
28            match m with 
29      [ O \Rightarrow false
30            | (S q) \Rightarrow eqb p q]].
31            
32 theorem eqb_to_Prop: \forall n,m:nat. 
33 match (eqb n m) with
34 [ true  \Rightarrow n = m 
35 | false \Rightarrow n \neq m].
36 intros.
37 apply (nat_elim2
38 (\lambda n,m:nat.match (eqb n m) with
39 [ true  \Rightarrow n = m 
40 | false \Rightarrow n \neq m])).
41 intro.elim n1.
42 simplify.reflexivity.
43 simplify.apply not_eq_O_S.
44 intro.
45 simplify.unfold Not.
46 intro. apply (not_eq_O_S n1).apply sym_eq.assumption.
47 intros.simplify.
48 generalize in match H.
49 elim ((eqb n1 m1)).
50 simplify.apply eq_f.apply H1.
51 simplify.unfold Not.intro.apply H1.apply inj_S.assumption.
52 qed.
53
54 theorem eqb_elim : \forall n,m:nat.\forall P:bool \to Prop.
55 (n=m \to (P true)) \to (n \neq m \to (P false)) \to (P (eqb n m)). 
56 intros.
57 cut 
58 (match (eqb n m) with
59 [ true  \Rightarrow n = m
60 | false \Rightarrow n \neq m] \to (P (eqb n m))).
61 apply Hcut.apply eqb_to_Prop.
62 elim (eqb n m).
63 apply ((H H2)).
64 apply ((H1 H2)).
65 qed.
66
67 theorem eqb_n_n: \forall n. eqb n n = true.
68 intro.elim n.simplify.reflexivity.
69 simplify.assumption.
70 qed.
71
72 theorem eqb_true_to_eq: \forall n,m:nat.
73 eqb n m = true \to n = m.
74 intros.
75 change with 
76 match true with
77 [ true  \Rightarrow n = m 
78 | false \Rightarrow n \neq m].
79 rewrite < H.
80 apply eqb_to_Prop. 
81 qed.
82
83 theorem eqb_false_to_not_eq: \forall n,m:nat.
84 eqb n m = false \to n \neq m.
85 intros.
86 change with 
87 match false with
88 [ true  \Rightarrow n = m 
89 | false \Rightarrow n \neq m].
90 rewrite < H.
91 apply eqb_to_Prop. 
92 qed.
93
94 theorem eq_to_eqb_true: \forall n,m:nat.
95 n = m \to eqb n m = true.
96 intros.apply (eqb_elim n m).
97 intros. reflexivity.
98 intros.apply False_ind.apply (H1 H).
99 qed.
100
101 theorem not_eq_to_eqb_false: \forall n,m:nat.
102 \lnot (n = m) \to eqb n m = false.
103 intros.apply (eqb_elim n m).
104 intros. apply False_ind.apply (H H1).
105 intros.reflexivity.
106 qed.
107
108 let rec leb n m \def 
109 match n with 
110     [ O \Rightarrow true
111     | (S p) \Rightarrow
112         match m with 
113         [ O \Rightarrow false
114         | (S q) \Rightarrow leb p q]].
115
116 theorem leb_elim: \forall n,m:nat. \forall P:bool \to Prop. 
117 (n \leq m \to (P true)) \to (n \nleq m \to (P false)) \to
118 P (leb n m).
119 apply nat_elim2; intros; simplify
120   [apply H.apply le_O_n
121   |apply H1.apply not_le_Sn_O.
122   |apply H;intros
123     [apply H1.apply le_S_S.assumption.
124     |apply H2.unfold Not.intros.apply H3.apply le_S_S_to_le.assumption
125     ]
126   ]
127 qed.
128
129 theorem leb_true_to_le:\forall n,m.
130 leb n m = true \to (n \le m).
131 intros 2.
132 apply leb_elim
133   [intros.assumption
134   |intros.destruct H1.
135   ]
136 qed.
137
138 theorem leb_false_to_not_le:\forall n,m.
139 leb n m = false \to \lnot (n \le m).
140 intros 2.
141 apply leb_elim
142   [intros.destruct H1
143   |intros.assumption
144   ]
145 qed.
146 (*
147 theorem decidable_le: \forall n,m. n \leq m \lor n \nleq m. 
148 intros.
149 apply (leb_elim n m)
150   [intro.left.assumption
151   |intro.right.assumption
152   ]
153 qed.
154 *)
155
156 theorem le_to_leb_true: \forall n,m. n \leq m \to leb n m = true.
157 intros.apply leb_elim;intros
158   [reflexivity
159   |apply False_ind.apply H1.apply H.
160   ]
161 qed.
162
163 theorem lt_to_leb_false: \forall n,m. m < n \to leb n m = false.
164 intros.apply leb_elim;intros
165   [apply False_ind.apply (le_to_not_lt ? ? H1). assumption
166   |reflexivity
167   ]
168 qed.
169
170 theorem leb_to_Prop: \forall n,m:nat. 
171 match (leb n m) with
172 [ true  \Rightarrow n \leq m 
173 | false \Rightarrow n \nleq m].
174 apply nat_elim2;simplify
175   [exact le_O_n
176   |exact not_le_Sn_O
177   |intros 2.simplify.
178    elim ((leb n m));simplify
179     [apply le_S_S.apply H
180     |unfold Not.intros.apply H.apply le_S_S_to_le.assumption
181     ]
182   ]
183 qed.
184
185 (*
186 theorem leb_elim: \forall n,m:nat. \forall P:bool \to Prop. 
187 (n \leq m \to (P true)) \to (n \nleq m \to (P false)) \to
188 P (leb n m).
189 intros.
190 cut 
191 (match (leb n m) with
192 [ true  \Rightarrow n \leq m
193 | false \Rightarrow n \nleq m] \to (P (leb n m))).
194 apply Hcut.apply leb_to_Prop.
195 elim (leb n m).
196 apply ((H H2)).
197 apply ((H1 H2)).
198 qed.
199 *)
200
201 definition ltb ≝λn,m. leb n m ∧ notb (eqb n m).
202
203 theorem ltb_to_Prop :
204  ∀n,m.
205   match ltb n m with
206   [ true ⇒ n < m
207   | false ⇒ n ≮ m
208   ].
209 intros;
210 unfold ltb;
211 apply leb_elim;
212 apply eqb_elim;
213 intros;
214 simplify;
215 [ rewrite < H;
216   apply le_to_not_lt;
217   constructor 1
218 | apply (not_eq_to_le_to_lt ? ? H H1)
219 | rewrite < H;
220   apply le_to_not_lt;
221   constructor 1
222 | apply le_to_not_lt;
223   generalize in match (not_le_to_lt ? ? H1);
224   clear H1;
225   intro;
226   apply lt_to_le;
227   assumption
228 ].
229 qed.
230
231 theorem ltb_elim: ∀n,m:nat. ∀P:bool → Prop.
232 (n < m → (P true)) → (n ≮ m → (P false)) →
233 P (ltb n m).
234 intros.
235 cut
236 (match (ltb n m) with
237 [ true  ⇒ n < m
238 | false ⇒ n ≮ m] → (P (ltb n m))).
239 apply Hcut.apply ltb_to_Prop.
240 elim (ltb n m).
241 apply ((H H2)).
242 apply ((H1 H2)).
243 qed.
244
245 let rec nat_compare n m: compare \def
246 match n with
247 [ O \Rightarrow 
248     match m with 
249       [ O \Rightarrow EQ
250       | (S q) \Rightarrow LT ]
251 | (S p) \Rightarrow 
252     match m with 
253       [ O \Rightarrow GT
254       | (S q) \Rightarrow nat_compare p q]].
255
256 theorem nat_compare_n_n: \forall n:nat. nat_compare n n = EQ.
257 intro.elim n.
258 simplify.reflexivity.
259 simplify.assumption.
260 qed.
261
262 theorem nat_compare_S_S: \forall n,m:nat. 
263 nat_compare n m = nat_compare (S n) (S m).
264 intros.simplify.reflexivity.
265 qed.
266
267 theorem nat_compare_pred_pred: 
268 \forall n,m:nat.lt O n \to lt O m \to 
269 eq compare (nat_compare n m) (nat_compare (pred n) (pred m)).
270 intros.
271 apply (lt_O_n_elim n H).
272 apply (lt_O_n_elim m H1).
273 intros.
274 simplify.reflexivity.
275 qed.
276
277 theorem nat_compare_to_Prop: \forall n,m:nat. 
278 match (nat_compare n m) with
279   [ LT \Rightarrow n < m
280   | EQ \Rightarrow n=m
281   | GT \Rightarrow m < n ].
282 intros.
283 apply (nat_elim2 (\lambda n,m.match (nat_compare n m) with
284   [ LT \Rightarrow n < m
285   | EQ \Rightarrow n=m
286   | GT \Rightarrow m < n ])).
287 intro.elim n1.simplify.reflexivity.
288 simplify.unfold lt.apply le_S_S.apply le_O_n.
289 intro.simplify.unfold lt.apply le_S_S. apply le_O_n.
290 intros 2.simplify.elim ((nat_compare n1 m1)).
291 simplify. unfold lt. apply le_S_S.apply H.
292 simplify. apply eq_f. apply H.
293 simplify. unfold lt.apply le_S_S.apply H.
294 qed.
295
296 theorem nat_compare_n_m_m_n: \forall n,m:nat. 
297 nat_compare n m = compare_invert (nat_compare m n).
298 intros. 
299 apply (nat_elim2 (\lambda n,m. nat_compare n m = compare_invert (nat_compare m n))).
300 intros.elim n1.simplify.reflexivity.
301 simplify.reflexivity.
302 intro.elim n1.simplify.reflexivity.
303 simplify.reflexivity.
304 intros.simplify.elim H.reflexivity.
305 qed.
306      
307 theorem nat_compare_elim : \forall n,m:nat. \forall P:compare \to Prop.
308 (n < m \to P LT) \to (n=m \to P EQ) \to (m < n \to P GT) \to 
309 (P (nat_compare n m)).
310 intros.
311 cut (match (nat_compare n m) with
312 [ LT \Rightarrow n < m
313 | EQ \Rightarrow n=m
314 | GT \Rightarrow m < n] \to
315 (P (nat_compare n m))).
316 apply Hcut.apply nat_compare_to_Prop.
317 elim ((nat_compare n m)).
318 apply ((H H3)).
319 apply ((H1 H3)).
320 apply ((H2 H3)).
321 qed.