]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/compare.ma
added homepage URL, now we have one
[helm.git] / helm / matita / library / nat / compare.ma
index 16d689e4cfafb19c34363ccdd6f18e8a6a71abc5..3a8e4a87ca50a7416e2df5f72920ea151977f41c 100644 (file)
@@ -28,13 +28,13 @@ match n with
        
 theorem leb_to_Prop: \forall n,m:nat. 
 match (leb n m) with
-[ true  \Rightarrow (le n m) 
-| false \Rightarrow (Not (le n m))].
+[ true  \Rightarrow n \leq m 
+| false \Rightarrow \lnot (n \leq 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))]).
+[ true  \Rightarrow n \leq m 
+| false \Rightarrow \lnot (n \leq m)]).
 simplify.exact le_O_n.
 simplify.exact not_le_Sn_O.
 intros 2.simplify.elim (leb n1 m1).
@@ -43,13 +43,13 @@ simplify.intros.apply H.apply le_S_S_to_le.assumption.
 qed.
 
 theorem leb_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
+(n \leq m \to (P true)) \to (\not (n \leq 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)).
+[ true  \Rightarrow n \leq m
+| false \Rightarrow \lnot (n \leq m)] \to (P (leb n m)).
 apply Hcut.apply leb_to_Prop.
 elim leb n m.
 apply (H H2).
@@ -78,16 +78,31 @@ nat_compare n m = nat_compare (S n) (S m).
 intros.simplify.reflexivity.
 qed.
 
+theorem S_pred: \forall n:nat.lt O n \to eq nat n (S (pred n)).
+intro.elim n.apply False_ind.exact not_le_Sn_O O H.
+apply eq_f.apply pred_Sn.
+qed.
+
+theorem nat_compare_pred_pred: 
+\forall n,m:nat.lt O n \to lt O m \to 
+eq compare (nat_compare n m) (nat_compare (pred n) (pred m)).
+intros.
+apply lt_O_n_elim n H.
+apply lt_O_n_elim m H1.
+intros.
+simplify.reflexivity.
+qed.
+
 theorem nat_compare_to_Prop: \forall n,m:nat. 
 match (nat_compare n m) with
-  [ LT \Rightarrow (lt n m)
+  [ LT \Rightarrow n < m
   | EQ \Rightarrow n=m
-  | GT \Rightarrow (lt m n) ].
+  | GT \Rightarrow m < n ].
 intros.
 apply nat_elim2 (\lambda n,m.match (nat_compare n m) with
-  [ LT \Rightarrow (lt n m)
+  [ LT \Rightarrow n < m
   | EQ \Rightarrow n=m
-  | GT \Rightarrow (lt m n) ]).
+  | GT \Rightarrow m < n ]).
 intro.elim n1.simplify.reflexivity.
 simplify.apply le_S_S.apply le_O_n.
 intro.simplify.apply le_S_S. apply le_O_n.
@@ -109,13 +124,13 @@ intros.simplify.elim H.reflexivity.
 qed.
      
 theorem nat_compare_elim : \forall n,m:nat. \forall P:compare \to Prop.
-((lt n m) \to (P LT)) \to (n=m \to (P EQ)) \to ((lt m n) \to (P GT)) \to 
+(n < m \to P LT) \to (n=m \to P EQ) \to (m < n \to P GT) \to 
 (P (nat_compare n m)).
 intros.
 cut match (nat_compare n m) with
-[ LT \Rightarrow (lt n m)
+[ LT \Rightarrow n < m
 | EQ \Rightarrow n=m
-| GT \Rightarrow (lt m n)] \to
+| GT \Rightarrow m < n] \to
 (P (nat_compare n m)).
 apply Hcut.apply nat_compare_to_Prop.
 elim (nat_compare n m).