X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Flibrary%2Fnat%2Fcompare.ma;h=2647315804661a23db188e36b4955c0b01c5152a;hb=4167cea65ca58897d1a3dbb81ff95de5074700cc;hp=852cd7b6551f44091425aebcbac65ae7d314c6fb;hpb=39eee3ee34bd22321d7b80c2685395345a0957e0;p=helm.git diff --git a/helm/matita/library/nat/compare.ma b/helm/matita/library/nat/compare.ma index 852cd7b65..264731580 100644 --- a/helm/matita/library/nat/compare.ma +++ b/helm/matita/library/nat/compare.ma @@ -32,36 +32,77 @@ 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 +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. +simplify.unfold Not. +intro. apply (not_eq_O_S n1).apply sym_eq.assumption. intros.simplify. generalize in match H. -elim (eqb n1 m1). +elim ((eqb n1 m1)). simplify.apply eq_f.apply H1. -simplify.intro.apply H1.apply inj_S.assumption. +simplify.unfold Not.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 +(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). +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 eqb_true_to_eq: \forall n,m:nat. +eqb n m = true \to n = m. +intros. +change with +match true with +[ true \Rightarrow n = m +| false \Rightarrow n \neq m]. +rewrite < H. +apply eqb_to_Prop. +qed. + +theorem eqb_false_to_not_eq: \forall n,m:nat. +eqb n m = false \to n \neq m. +intros. +change with +match false with +[ true \Rightarrow n = m +| false \Rightarrow n \neq m]. +rewrite < H. +apply eqb_to_Prop. +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 @@ -75,31 +116,31 @@ 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 +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). +intros 2.simplify.elim ((leb n1 m1)). simplify.apply le_S_S.apply H. -simplify.intros.apply H.apply le_S_S_to_le.assumption. +simplify.unfold Not.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 +(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). -apply (H1 H2). +elim (leb n m). +apply ((H H2)). +apply ((H1 H2)). qed. let rec nat_compare n m: compare \def @@ -125,7 +166,7 @@ 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. +intro.elim n.apply False_ind.exact (not_le_Sn_O O H). apply eq_f.apply pred_Sn. qed. @@ -133,8 +174,8 @@ 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. +apply (lt_O_n_elim n H). +apply (lt_O_n_elim m H1). intros. simplify.reflexivity. qed. @@ -145,23 +186,23 @@ match (nat_compare n m) with | EQ \Rightarrow n=m | GT \Rightarrow m < n ]. intros. -apply nat_elim2 (\lambda n,m.match (nat_compare n m) with +apply (nat_elim2 (\lambda n,m.match (nat_compare n m) with [ LT \Rightarrow n < m | EQ \Rightarrow n=m - | GT \Rightarrow 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. -intros 2.simplify.elim (nat_compare n1 m1). -simplify. apply le_S_S.apply H. -simplify. apply le_S_S.apply H. +simplify.unfold lt.apply le_S_S.apply le_O_n. +intro.simplify.unfold lt.apply le_S_S. apply le_O_n. +intros 2.simplify.elim ((nat_compare n1 m1)). +simplify. unfold lt. apply le_S_S.apply H. simplify. apply eq_f. apply H. +simplify. unfold lt.apply le_S_S.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)). +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. @@ -173,14 +214,14 @@ 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 +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)). +(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). +elim ((nat_compare n m)). +apply ((H H3)). +apply ((H1 H3)). +apply ((H2 H3)). qed.