]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/compare.ma
New version of the library, a bit more structured.
[helm.git] / helm / matita / library / nat / compare.ma
diff --git a/helm/matita/library/nat/compare.ma b/helm/matita/library/nat/compare.ma
new file mode 100644 (file)
index 0000000..af3ca38
--- /dev/null
@@ -0,0 +1,59 @@
+(**************************************************************************)
+(*       ___                                                               *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||       A.Asperti, C.Sacerdoti Coen,                          *)
+(*      ||A||       E.Tassi, S.Zacchiroli                                 *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU Lesser General Public License Version 2.1         *)
+(*                                                                        *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/nat/compare.ma".
+
+include "nat/orders.ma".
+include "datatypes/bool.ma".
+
+let rec leb n m \def 
+match n with 
+    [ O \Rightarrow true
+    | (S p) \Rightarrow
+       match m with 
+        [ O \Rightarrow false
+       | (S q) \Rightarrow leb p q]].
+       
+theorem leb_to_Prop: \forall n,m:nat. 
+match (leb n m) with
+[ true  \Rightarrow (le n m) 
+| false \Rightarrow (Not (le n m))].
+intros.
+apply nat_elim2
+(\lambda n,m:nat.match (leb n m) with
+[ true  \Rightarrow (le n m) 
+| false \Rightarrow (Not (le n m))]).
+simplify.exact le_O_n.
+simplify.exact not_le_Sn_O.
+intros 2.simplify.elim (leb n1 m1).
+simplify.apply le_S_S.apply H.
+simplify.intros.apply H.apply le_S_S_to_le.assumption.
+qed.
+
+theorem le_elim: \forall n,m:nat. \forall P:bool \to Prop. 
+((le n m) \to (P true)) \to ((Not (le n m)) \to (P false)) \to
+P (leb n m).
+intros.
+cut 
+match (leb n m) with
+[ true  \Rightarrow (le n m) 
+| false \Rightarrow (Not (le n m))] \to (P (leb n m)).
+apply Hcut.apply leb_to_Prop.
+elim leb n m.
+apply (H H2).
+apply (H1 H2).
+qed.
+
+
+