X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Flibrary%2Fnat%2Fcompare.ma;h=fec46ae2fe1f65f8106020bc4a038553beb1667b;hb=dc861d214cb992a898f81752614201b8074eef12;hp=af3ca38f83c0cd02478e8c85fc118e82d906a0ec;hpb=244d65f63ca6a736b871f9f91328fe8c5524ff05;p=helm.git diff --git a/helm/matita/library/nat/compare.ma b/helm/matita/library/nat/compare.ma index af3ca38f8..fec46ae2f 100644 --- a/helm/matita/library/nat/compare.ma +++ b/helm/matita/library/nat/compare.ma @@ -12,10 +12,11 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/nat/compare.ma". +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 @@ -41,7 +42,7 @@ 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. +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 P (leb n m). intros. @@ -55,5 +56,70 @@ 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.(eq compare (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)). +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) ]. +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) ]). +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. +eq compare (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))). +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. +((lt n m) \to (P LT)) \to ((eq nat n m) \to (P EQ)) \to ((lt 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 +(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.