1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/library_auto/nat/compare".
17 include "datatypes/bool.ma".
18 include "datatypes/compare.ma".
19 include "auto/nat/orders.ma".
26 | (S q) \Rightarrow false]
30 | (S q) \Rightarrow eqb p q]].
32 theorem eqb_to_Prop: \forall n,m:nat.
34 [ true \Rightarrow n = m
35 | false \Rightarrow n \neq m].
38 (\lambda n,m:nat.match (eqb n m) with
39 [ true \Rightarrow n = m
40 | false \Rightarrow n \neq m]))
52 apply (not_eq_O_S n1).
58 generalize in match H.
59 elim ((eqb n1 m1));simplify
72 theorem eqb_elim : \forall n,m:nat.\forall P:bool \to Prop.
73 (n=m \to (P true)) \to (n \neq m \to (P false)) \to (P (eqb n m)).
77 [ true \Rightarrow n = m
78 | false \Rightarrow n \neq m] \to (P (eqb n m)))
80 (* qui auto non conclude il goal*)
83 [ (*qui auto non conclude il goal*)
85 | (*qui auto non conclude il goal*)
91 theorem eqb_n_n: \forall n. eqb n n = true.
94 (*[ simplify.reflexivity
95 | simplify.assumption.
99 theorem eqb_true_to_eq: \forall n,m:nat.
100 eqb n m = true \to n = m.
104 [ true \Rightarrow n = m
105 | false \Rightarrow n \neq m].
107 (*qui auto non conclude il goal*)
111 theorem eqb_false_to_not_eq: \forall n,m:nat.
112 eqb n m = false \to n \neq m.
116 [ true \Rightarrow n = m
117 | false \Rightarrow n \neq m].
119 (*qui auto non conclude il goal*)
123 theorem eq_to_eqb_true: \forall n,m:nat.
124 n = m \to eqb n m = true.
127 (*apply (eqb_elim n m)
128 [ intros. reflexivity
129 | intros.apply False_ind.apply (H1 H)
133 theorem not_eq_to_eqb_false: \forall n,m:nat.
134 \lnot (n = m) \to eqb n m = false.
135 intros.apply (eqb_elim n m);intros
147 [ O \Rightarrow false
148 | (S q) \Rightarrow leb p q]].
150 theorem leb_to_Prop: \forall n,m:nat.
152 [ true \Rightarrow n \leq m
153 | false \Rightarrow n \nleq m].
156 (\lambda n,m:nat.match (leb n m) with
157 [ true \Rightarrow n \leq m
158 | false \Rightarrow n \nleq m]))
165 elim ((leb n1 m1));simplify
167 (*qui auto non conclude il goal*)
173 (*apply le_S_S_to_le.
179 theorem leb_elim: \forall n,m:nat. \forall P:bool \to Prop.
180 (n \leq m \to (P true)) \to (n \nleq m \to (P false)) \to
184 (match (leb n m) with
185 [ true \Rightarrow n \leq m
186 | false \Rightarrow n \nleq m] \to (P (leb n m)))
188 (*qui auto non conclude il goal*)
191 [ (*qui auto non conclude il goal*)
193 | (*qui auto non conclude il goal*)
199 let rec nat_compare n m: compare \def
204 | (S q) \Rightarrow LT ]
208 | (S q) \Rightarrow nat_compare p q]].
210 theorem nat_compare_n_n: \forall n:nat. nat_compare n n = EQ.
220 theorem nat_compare_S_S: \forall n,m:nat.
221 nat_compare n m = nat_compare (S n) (S m).
223 (*simplify.reflexivity.*)
226 theorem S_pred: \forall n:nat.lt O n \to eq nat n (S (pred n)).
230 exact (not_le_Sn_O O H)
236 theorem nat_compare_pred_pred:
237 \forall n,m:nat.lt O n \to lt O m \to
238 eq compare (nat_compare n m) (nat_compare (pred n) (pred m)).
240 apply (lt_O_n_elim n H).
241 apply (lt_O_n_elim m H1).
244 (*simplify.reflexivity.*)
247 theorem nat_compare_to_Prop: \forall n,m:nat.
248 match (nat_compare n m) with
249 [ LT \Rightarrow n < m
251 | GT \Rightarrow m < n ].
253 apply (nat_elim2 (\lambda n,m.match (nat_compare n m) with
254 [ LT \Rightarrow n < m
256 | GT \Rightarrow m < n ]))
258 elim n1;simplify;auto
272 elim ((nat_compare n1 m1));simplify
275 (*qui auto non chiude il goal*)
278 (*qui auto non chiude il goal*)
282 (*qui auto non chiude il goal*)
288 theorem nat_compare_n_m_m_n: \forall n,m:nat.
289 nat_compare n m = compare_invert (nat_compare m n).
291 apply (nat_elim2 (\lambda n,m. nat_compare n m = compare_invert (nat_compare m n)));intros
292 [ elim n1;auto(*;simplify;reflexivity*)
293 | elim n1;auto(*;simplify;reflexivity*)
295 (*simplify.elim H.reflexivity*)
299 theorem nat_compare_elim : \forall n,m:nat. \forall P:compare \to Prop.
300 (n < m \to P LT) \to (n=m \to P EQ) \to (m < n \to P GT) \to
301 (P (nat_compare n m)).
303 cut (match (nat_compare n m) with
304 [ LT \Rightarrow n < m
306 | GT \Rightarrow m < n] \to
307 (P (nat_compare n m)))
309 (*auto non chiude il goal*)
310 apply nat_compare_to_Prop
311 | elim ((nat_compare n m))
312 [ (*auto non chiude il goal*)
314 | (*auto non chiude il goal*)
316 | (*auto non chiude il goal*)