]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/library/nat/compare.ma
Towards chebyshev.
[helm.git] / matita / library / nat / compare.ma
index 2647315804661a23db188e36b4955c0b01c5152a..b002d78cccdedd6ad690ffeea67aaa0402c14288 100644 (file)
@@ -112,23 +112,77 @@ match n with
        match m with 
         [ O \Rightarrow false
        | (S q) \Rightarrow leb p q]].
-       
+
+theorem leb_elim: \forall n,m:nat. \forall P:bool \to Prop. 
+(n \leq m \to (P true)) \to (n \nleq m \to (P false)) \to
+P (leb n m).
+apply nat_elim2; intros; simplify
+  [apply H.apply le_O_n
+  |apply H1.apply not_le_Sn_O.
+  |apply H;intros
+    [apply H1.apply le_S_S.assumption.
+    |apply H2.unfold Not.intros.apply H3.apply le_S_S_to_le.assumption
+    ]
+  ]
+qed.
+
+theorem leb_true_to_le:\forall n,m.
+leb n m = true \to (n \le m).
+intros 2.
+apply leb_elim
+  [intros.assumption
+  |intros.destruct H1.
+  ]
+qed.
+
+theorem leb_false_to_not_le:\forall n,m.
+leb n m = false \to \lnot (n \le m).
+intros 2.
+apply leb_elim
+  [intros.destruct H1
+  |intros.assumption
+  ]
+qed.
+(*
+theorem decidable_le: \forall n,m. n \leq m \lor n \nleq m. 
+intros.
+apply (leb_elim n m)
+  [intro.left.assumption
+  |intro.right.assumption
+  ]
+qed.
+*)
+
+theorem le_to_leb_true: \forall n,m. n \leq m \to leb n m = true.
+intros.apply leb_elim;intros
+  [reflexivity
+  |apply False_ind.apply H1.apply H.
+  ]
+qed.
+
+theorem lt_to_leb_false: \forall n,m. m < n \to leb n m = false.
+intros.apply leb_elim;intros
+  [apply False_ind.apply (le_to_not_lt ? ? H1). assumption
+  |reflexivity
+  ]
+qed.
+
 theorem leb_to_Prop: \forall n,m:nat. 
 match (leb n m) with
 [ 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 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)).
-simplify.apply le_S_S.apply H.
-simplify.unfold Not.intros.apply H.apply le_S_S_to_le.assumption.
+apply nat_elim2;simplify
+  [exact le_O_n
+  |exact not_le_Sn_O
+  |intros 2.simplify.
+   elim ((leb n m));simplify
+    [apply le_S_S.apply H
+    |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 (n \nleq m \to (P false)) \to
 P (leb n m).
@@ -142,6 +196,51 @@ elim (leb n m).
 apply ((H H2)).
 apply ((H1 H2)).
 qed.
+*)
+
+definition ltb ≝λn,m. leb n m ∧ notb (eqb n m).
+
+theorem ltb_to_Prop :
+ ∀n,m.
+  match ltb n m with
+  [ true ⇒ n < m
+  | false ⇒ n ≮ m
+  ].
+intros;
+unfold ltb;
+apply leb_elim;
+apply eqb_elim;
+intros;
+simplify;
+[ rewrite < H;
+  apply le_to_not_lt;
+  constructor 1
+| apply (not_eq_to_le_to_lt ? ? H H1)
+| rewrite < H;
+  apply le_to_not_lt;
+  constructor 1
+| apply le_to_not_lt;
+  generalize in match (not_le_to_lt ? ? H1);
+  clear H1;
+  intro;
+  apply lt_to_le;
+  assumption
+].
+qed.
+
+theorem ltb_elim: ∀n,m:nat. ∀P:bool → Prop.
+(n < m → (P true)) → (n ≮ m → (P false)) →
+P (ltb n m).
+intros.
+cut
+(match (ltb n m) with
+[ true  ⇒ n < m
+| false ⇒ n ≮ m] → (P (ltb n m))).
+apply Hcut.apply ltb_to_Prop.
+elim (ltb n m).
+apply ((H H2)).
+apply ((H1 H2)).
+qed.
 
 let rec nat_compare n m: compare \def
 match n with
@@ -165,11 +264,6 @@ 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)).