X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Flibrary%2Fnat%2Fcompare.ma;h=3a8e4a87ca50a7416e2df5f72920ea151977f41c;hb=b9af9f1c0de6a1735b492f5c793a87a8fce218cc;hp=d148dfd310b2b8be46fdb8a5e4b854e71d6ac34f;hpb=f263e4ec717d5ec2e7f9c057855f8223f81baae8;p=helm.git diff --git a/helm/matita/library/nat/compare.ma b/helm/matita/library/nat/compare.ma index d148dfd31..3a8e4a87c 100644 --- a/helm/matita/library/nat/compare.ma +++ b/helm/matita/library/nat/compare.ma @@ -16,6 +16,7 @@ set "baseuri" "cic:/matita/nat/compare". include "nat/orders.ma". include "datatypes/bool.ma". +include "datatypes/compare.ma". let rec leb n m \def match n with @@ -27,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). @@ -41,19 +42,99 @@ 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 +theorem leb_elim: \forall n,m:nat. \forall P:bool \to Prop. +(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). apply (H1 H2). qed. +let rec nat_compare n m: compare \def +match n with +[ O \Rightarrow + match m with + [ O \Rightarrow EQ + | (S q) \Rightarrow LT ] +| (S p) \Rightarrow + match m with + [ O \Rightarrow GT + | (S q) \Rightarrow nat_compare p q]]. +theorem nat_compare_n_n: \forall n:nat. nat_compare n n = EQ. +intro.elim n. +simplify.reflexivity. +simplify.assumption. +qed. + +theorem nat_compare_S_S: \forall n,m:nat. +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 n < m + | EQ \Rightarrow n=m + | GT \Rightarrow m < n ]. +intros. +apply nat_elim2 (\lambda n,m.match (nat_compare n m) with + [ LT \Rightarrow n < m + | EQ \Rightarrow n=m + | 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. +intros 2.simplify.elim (nat_compare n1 m1). +simplify. apply le_S_S.apply H. +simplify. apply le_S_S.apply H. +simplify. apply eq_f. apply H. +qed. + +theorem nat_compare_n_m_m_n: \forall n,m:nat. +nat_compare n m = compare_invert (nat_compare m n). +intros. +apply nat_elim2 (\lambda n,m. nat_compare n m = compare_invert (nat_compare m n)). +intros.elim n1.simplify.reflexivity. +simplify.reflexivity. +intro.elim n1.simplify.reflexivity. +simplify.reflexivity. +intros.simplify.elim H.reflexivity. +qed. + +theorem nat_compare_elim : \forall n,m:nat. \forall P:compare \to Prop. +(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 n < m +| EQ \Rightarrow n=m +| GT \Rightarrow m < n] \to +(P (nat_compare n m)). +apply Hcut.apply nat_compare_to_Prop. +elim (nat_compare n m). +apply (H H3). +apply (H2 H3). +apply (H1 H3). +qed.