]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/RELATIONAL/NLE/order.ma
removing old contribs
[helm.git] / helm / software / matita / contribs / RELATIONAL / NLE / order.ma
index c50e7befd15b761cb0ee55035fecb7f95a2daf38..473ac601d889d34fa47e01e64742d21ae1b7c3f1 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/RELATIONAL/NLE/order".
+
 
 include "NLE/inv.ma".
 
-theorem nle_refl: \forall x. x <= x.
+theorem nle_refl: ∀x. x ≤ x.
  intros; elim x; clear x; autobatch.
 qed.
 
-theorem nle_trans: \forall x,y. x <= y \to
-                   \forall z. y <= z \to x <= z.
- intros 3. elim H; clear H x y;
+theorem nle_trans: ∀x,y. x ≤ y → ∀z. y ≤ z → x ≤ z.
+ intros 3; elim H; clear H x y;
  [ autobatch
- | lapply linear nle_inv_succ_1 to H3. decompose. subst. 
+ | lapply linear nle_inv_succ_1 to H3. decompose. destruct. 
    autobatch
  ].
 qed.
 
-theorem nle_false: \forall x,y. x <= y \to y < x \to False.
+theorem nle_false: ∀x,y. x ≤ y → y < x → False.
  intros 3; elim H; clear H x y; autobatch.
 qed.
 
-theorem nle_irrefl: \forall x. x < x \to False.
+theorem nle_irrefl: ∀x. x < x → False.
  intros. autobatch.
 qed.
 
-theorem nle_irrefl_ei: \forall x, z. z <= x \to z = succ x \to False.
- intros 3. elim H; clear H x z; subst. autobatch.
+theorem nle_irrefl_ei: ∀x, z. z ≤ x → z = succ x → False.
+ intros 3; elim H; clear H x z; destruct; autobatch.
 qed.
 
-theorem nle_irrefl_smart: \forall x. x < x \to False.
+theorem nle_irrefl_smart: ∀x. x < x → False.
  intros 1. elim x; clear x; autobatch.
 qed.
 
-theorem nle_lt_or_eq: \forall y, x. x <= y \to x < y \lor x = y.
- intros. elim H; clear H x y;
+theorem nle_lt_or_eq: ∀y, x. x ≤ y → x < y ∨ x = y.
+ intros; elim H; clear H x y;
  [ elim n; clear n
  | decompose
  ]; autobatch.