X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Flibrary%2Fnat%2Fcompare.ma;h=55e526c8fad0f80eade4775141c99326975efee8;hb=8b55faddb06e3c4b0a13839210bb49170939b33e;hp=fec46ae2fe1f65f8106020bc4a038553beb1667b;hpb=b5ab291cf2ad1431128dd0b1ab629169a3cdeb50;p=helm.git diff --git a/helm/matita/library/nat/compare.ma b/helm/matita/library/nat/compare.ma index fec46ae2f..55e526c8f 100644 --- a/helm/matita/library/nat/compare.ma +++ b/helm/matita/library/nat/compare.ma @@ -1,5 +1,5 @@ (**************************************************************************) -(* ___ *) +(* ___ *) (* ||M|| *) (* ||A|| A project by Andrea Asperti *) (* ||T|| *) @@ -14,9 +14,74 @@ set "baseuri" "cic:/matita/nat/compare". -include "nat/orders.ma". include "datatypes/bool.ma". include "datatypes/compare.ma". +include "nat/orders.ma". + +let rec eqb n m \def +match n with + [ O \Rightarrow + match m with + [ O \Rightarrow true + | (S q) \Rightarrow false] + | (S p) \Rightarrow + match m with + [ O \Rightarrow false + | (S q) \Rightarrow eqb p q]]. + +theorem eqb_to_Prop: \forall n,m:nat. +match (eqb n m) with +[ true \Rightarrow n = m +| false \Rightarrow n \neq m]. +intros. +apply nat_elim2 +(\lambda n,m:nat.match (eqb n m) with +[ true \Rightarrow n = m +| false \Rightarrow n \neq m]). +intro.elim n1. +simplify.reflexivity. +simplify.apply not_eq_O_S. +intro. +simplify. +intro. apply not_eq_O_S n1.apply sym_eq.assumption. +intros.simplify. +generalize in match H. +elim (eqb n1 m1). +simplify.apply eq_f.apply H1. +simplify.intro.apply H1.apply inj_S.assumption. +qed. + +theorem eqb_elim : \forall n,m:nat.\forall P:bool \to Prop. +(n=m \to (P true)) \to (n \neq m \to (P false)) \to (P (eqb n m)). +intros. +cut +match (eqb n m) with +[ true \Rightarrow n = m +| false \Rightarrow n \neq m] \to (P (eqb n m)). +apply Hcut.apply eqb_to_Prop. +elim eqb n m. +apply (H H2). +apply (H1 H2). +qed. + +theorem eqb_n_n: \forall n. eqb n n = true. +intro.elim n.simplify.reflexivity. +simplify.assumption. +qed. + +theorem eq_to_eqb_true: \forall n,m:nat. +n = m \to eqb n m = true. +intros.apply eqb_elim n m. +intros. reflexivity. +intros.apply False_ind.apply H1 H. +qed. + +theorem not_eq_to_eqb_false: \forall n,m:nat. +\lnot (n = m) \to eqb n m = false. +intros.apply eqb_elim n m. +intros. apply False_ind.apply H H1. +intros.reflexivity. +qed. let rec leb n m \def match n with @@ -28,13 +93,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 n \nleq 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 n \nleq m]). simplify.exact le_O_n. simplify.exact not_le_Sn_O. intros 2.simplify.elim (leb n1 m1). @@ -43,13 +108,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 (n \nleq 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 n \nleq m] \to (P (leb n m)). apply Hcut.apply leb_to_Prop. elim leb n m. apply (H H2). @@ -67,40 +132,55 @@ match n with [ O \Rightarrow GT | (S q) \Rightarrow nat_compare p q]]. -theorem nat_compare_n_n: \forall n:nat.(eq compare (nat_compare n n) EQ). +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. -eq compare (nat_compare n m) (nat_compare (S n) (S m)). +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) - | EQ \Rightarrow (eq nat n m) - | GT \Rightarrow (lt m n) ]. + [ 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 (lt n m) - | EQ \Rightarrow (eq nat n m) - | GT \Rightarrow (lt m n) ]). + [ 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. +simplify. apply le_S_S.apply H. qed. theorem nat_compare_n_m_m_n: \forall n,m:nat. -eq compare (nat_compare n m) (compare_invert (nat_compare m n)). +nat_compare n m = compare_invert (nat_compare m n). intros. -apply nat_elim2 (\lambda n,m.eq compare (nat_compare n m) (compare_invert (nat_compare m n))). +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. @@ -109,17 +189,17 @@ 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 ((eq nat 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) -| EQ \Rightarrow (eq nat n m) -| GT \Rightarrow (lt m n)] \to +[ 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). +apply (H2 H3). qed.