X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Flibrary%2Fnat%2Fcompare.ma;h=55e526c8fad0f80eade4775141c99326975efee8;hb=8b55faddb06e3c4b0a13839210bb49170939b33e;hp=194b38d84f255e3924a3273919f01bb96d320ca5;hpb=03fcee16d9c262aad38a47d0a409b684a965cc3f;p=helm.git diff --git a/helm/matita/library/nat/compare.ma b/helm/matita/library/nat/compare.ma index 194b38d84..55e526c8f 100644 --- a/helm/matita/library/nat/compare.ma +++ b/helm/matita/library/nat/compare.ma @@ -14,9 +14,9 @@ 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 @@ -32,18 +32,18 @@ match n with theorem eqb_to_Prop: \forall n,m:nat. match (eqb n m) with [ true \Rightarrow n = m -| false \Rightarrow \lnot (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 \lnot (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. +intro. apply not_eq_O_S n1.apply sym_eq.assumption. intros.simplify. generalize in match H. elim (eqb n1 m1). @@ -52,18 +52,37 @@ 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 (\lnot n=m \to (P false)) \to (P (eqb n m)). +(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 \lnot (n = m)] \to (P (eqb 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 [ O \Rightarrow true @@ -75,12 +94,12 @@ match n with theorem leb_to_Prop: \forall n,m:nat. match (leb n m) with [ true \Rightarrow n \leq m -| false \Rightarrow \lnot (n \leq m)]. +| false \Rightarrow n \nleq m]. intros. apply nat_elim2 (\lambda n,m:nat.match (leb n m) with [ true \Rightarrow n \leq m -| false \Rightarrow \lnot (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). @@ -89,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. -(n \leq m \to (P true)) \to (\not (n \leq 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 n \leq m -| false \Rightarrow \lnot (n \leq m)] \to (P (leb n m)). +| false \Rightarrow n \nleq m] \to (P (leb n m)). apply Hcut.apply leb_to_Prop. elim leb n m. apply (H H2). @@ -154,8 +173,8 @@ 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. @@ -181,6 +200,6 @@ cut match (nat_compare n m) with 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.