]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/orders.ma
New version of the library, a bit more structured.
[helm.git] / helm / matita / library / nat / orders.ma
diff --git a/helm/matita/library/nat/orders.ma b/helm/matita/library/nat/orders.ma
new file mode 100644 (file)
index 0000000..4c1b1a0
--- /dev/null
@@ -0,0 +1,158 @@
+(**************************************************************************)
+(*       ___                                                               *)
+(*      ||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/orders.ma".
+
+include "logic/connectives.ma".
+include "logic/equality.ma".
+include "nat/nat.ma".
+include "nat/plus.ma".
+include "higher_order_defs/ordering.ma".
+
+(* definitions *)
+inductive le (n:nat) : nat \to Prop \def
+  | le_n : le n n
+  | le_S : \forall m:nat. le n m \to le n (S m).
+
+definition lt: nat \to nat \to Prop \def
+\lambda n,m:nat.le (S n) m.
+
+definition ge: nat \to nat \to Prop \def
+\lambda n,m:nat.le m n.
+
+definition gt: nat \to nat \to Prop \def
+\lambda n,m:nat.lt m n.
+
+theorem transitive_le : transitive nat le.
+simplify.intros.elim H1.
+assumption.
+apply le_S.assumption.
+qed.
+
+theorem trans_le: \forall n,m,p:nat. le n m \to le m p \to le n p
+\def transitive_le.
+
+theorem le_S_S: \forall n,m:nat. le n m \to le (S n) (S m).
+intros.elim H.
+apply le_n.
+apply le_S.assumption.
+qed.
+
+theorem le_O_n : \forall n:nat. le O n.
+intros.elim n.
+apply le_n.apply 
+le_S. assumption.
+qed.
+
+theorem le_n_Sn : \forall n:nat. le n (S n).
+intros. apply le_S.apply le_n.
+qed.
+
+theorem le_pred_n : \forall n:nat. le (pred n) n.
+intros.elim n.
+simplify.apply le_n.
+simplify.apply le_n_Sn.
+qed.
+
+theorem le_S_S_to_le : \forall n,m:nat. le (S n) (S m) \to le n m.
+intros.change with le (pred (S n)) (pred (S m)).
+elim H.apply le_n.apply trans_le ? (pred x).assumption.
+apply le_pred_n.
+qed.
+
+theorem leS_to_not_zero : \forall n,m:nat. (le (S n) m ) \to not_zero m.
+intros.elim H.exact I.exact I.
+qed.
+
+(* not le *)
+theorem not_le_Sn_O: \forall n:nat. Not (le (S n) O).
+intros.simplify.intros.apply leS_to_not_zero ? ? H.
+qed.
+
+theorem not_le_Sn_n: \forall n:nat. Not (le (S n) n).
+intros.elim n.apply not_le_Sn_O.simplify.intros.cut le (S e1) e1.
+apply H.assumption.
+apply le_S_S_to_le.assumption.
+qed.
+
+(* le vs. lt *)
+theorem lt_to_le : \forall n,m:nat. lt n m \to le n m.
+simplify.intros.elim H.
+apply le_S. apply le_n.
+apply le_S. assumption.
+qed.
+
+theorem not_le_to_lt: \forall n,m:nat. Not (le n m) \to lt m n.
+intros 2.
+apply nat_elim2 (\lambda n,m.Not (le n m) \to lt m n).
+intros.apply absurd (le O n1).apply le_O_n.assumption.
+simplify.intros.apply le_S_S.apply le_O_n.
+simplify.intros.apply le_S_S.apply H.intros.apply H1.apply le_S_S.
+assumption.
+qed.
+
+theorem lt_to_not_le: \forall n,m:nat. lt n m \to Not (le m n).
+simplify.intros 3.elim H.
+apply not_le_Sn_n n H1.
+apply H2.apply lt_to_le. apply H3.
+qed.
+
+(* le elimination *)
+theorem le_n_O_to_eq : \forall n:nat. (le n O) \to (eq nat O n).
+intro.elim n.reflexivity.
+apply False_ind.
+(* non si applica not_le_Sn_O *)
+apply  (not_le_Sn_O ? H1).
+qed.
+
+theorem le_n_O_elim: \forall n:nat.le n O \to \forall P: nat \to Prop.
+P O \to P n.
+intro.elim n.
+assumption.
+apply False_ind.
+apply  (not_le_Sn_O ? H1).
+qed.
+
+theorem le_n_Sm_elim : \forall n,m:nat.le n (S m) \to 
+\forall P:Prop. (le (S n) (S m) \to P) \to (eq nat n (S m) \to P) \to P.
+intros 4.elim H.
+apply H2.reflexivity.
+apply H3. apply le_S_S. assumption.
+qed.
+
+(* other abstract properties *)
+theorem antisymmetric_le : antisymmetric nat le.
+simplify.intros 2.
+apply nat_elim2 (\lambda n,m.((le n m) \to (le m n) \to eq nat n m)).
+intros.apply le_n_O_to_eq.assumption.
+intros.apply False_ind.apply not_le_Sn_O ? H.
+intros.apply eq_f.apply H.
+apply le_S_S_to_le.assumption.
+apply le_S_S_to_le.assumption.
+qed.
+
+theorem antisym_le: \forall n,m:nat. le n m \to le m n \to eq nat n m
+\def antisymmetric_le.
+
+theorem decidable_le: \forall n,m:nat. decidable (le n m).
+intros.
+apply nat_elim2 (\lambda n,m.decidable (le n m)).
+intros.simplify.left.apply le_O_n.
+intros.simplify.right.exact not_le_Sn_O n1.
+intros 2.simplify.intro.elim H.
+left.apply le_S_S.assumption.
+right.intro.apply H1.apply le_S_S_to_le.assumption.
+qed.
+
+