]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/orders.ma
A little bit more of notation here and there.
[helm.git] / helm / matita / library / nat / orders.ma
index 3f788c6506c11ebd62f638c02e2f1cd68c031783..5389418b5237d7bde8cf4a9424bba777a5b38a62 100644 (file)
@@ -22,14 +22,23 @@ inductive le (n:nat) : nat \to Prop \def
   | le_n : le n n
   | le_S : \forall m:nat. le n m \to le n (S m).
 
+interpretation "natural 'less or equal to'" 'leq x y = (cic:/matita/nat/orders/le.ind#xpointer(1/1) x y).
+alias symbol "leq" (instance 0) = "natural 'less or equal to'".
+
 definition lt: nat \to nat \to Prop \def
-\lambda n,m:nat.le (S n) m.
+\lambda n,m:nat.(S n) \leq m.
+
+interpretation "natural 'less than'" 'lt x y = (cic:/matita/nat/orders/lt.con x y).
 
 definition ge: nat \to nat \to Prop \def
-\lambda n,m:nat.le m n.
+\lambda n,m:nat.m \leq n.
+
+interpretation "natural 'greater or equal to'" 'geq x y = (cic:/matita/nat/orders/ge.con x y).
 
 definition gt: nat \to nat \to Prop \def
-\lambda n,m:nat.lt m n.
+\lambda n,m:nat.m<n.
+
+interpretation "natural 'greater than'" 'gt x y = (cic:/matita/nat/orders/gt.con x y).
 
 theorem transitive_le : transitive nat le.
 simplify.intros.elim H1.
@@ -37,89 +46,89 @@ assumption.
 apply le_S.assumption.
 qed.
 
-theorem trans_le: \forall n,m,p:nat. le n m \to le m p \to le n p
+theorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p
 \def transitive_le.
 
-theorem le_S_S: \forall n,m:nat. le n m \to le (S n) (S m).
+theorem le_S_S: \forall n,m:nat. n \leq m \to S n \leq S m.
 intros.elim H.
 apply le_n.
 apply le_S.assumption.
 qed.
 
-theorem le_O_n : \forall n:nat. le O n.
+theorem le_O_n : \forall n:nat. O \leq n.
 intros.elim n.
 apply le_n.apply 
 le_S. assumption.
 qed.
 
-theorem le_n_Sn : \forall n:nat. le n (S n).
+theorem le_n_Sn : \forall n:nat. n \leq S n.
 intros. apply le_S.apply le_n.
 qed.
 
-theorem le_pred_n : \forall n:nat. le (pred n) n.
+theorem le_pred_n : \forall n:nat. pred n \leq n.
 intros.elim n.
 simplify.apply le_n.
 simplify.apply le_n_Sn.
 qed.
 
-theorem le_S_S_to_le : \forall n,m:nat. le (S n) (S m) \to le n m.
-intros.change with le (pred (S n)) (pred (S m)).
+theorem le_S_S_to_le : \forall n,m:nat. S n \leq S m \to n \leq m.
+intros.change with pred (S n) \leq pred (S m).
 elim H.apply le_n.apply trans_le ? (pred n1).assumption.
 apply le_pred_n.
 qed.
 
-theorem leS_to_not_zero : \forall n,m:nat. (le (S n) m ) \to not_zero m.
+theorem leS_to_not_zero : \forall n,m:nat. S n \leq m \to not_zero m.
 intros.elim H.exact I.exact I.
 qed.
 
 (* not le *)
-theorem not_le_Sn_O: \forall n:nat. Not (le (S n) O).
+theorem not_le_Sn_O: \forall n:nat. \lnot (S n \leq O).
 intros.simplify.intros.apply leS_to_not_zero ? ? H.
 qed.
 
-theorem not_le_Sn_n: \forall n:nat. Not (le (S n) n).
-intros.elim n.apply not_le_Sn_O.simplify.intros.cut le (S n1) n1.
+theorem not_le_Sn_n: \forall n:nat. \lnot (S n \leq n).
+intros.elim n.apply not_le_Sn_O.simplify.intros.cut S n1 \leq n1.
 apply H.assumption.
 apply le_S_S_to_le.assumption.
 qed.
 
 (* ??? this needs not le *)
-theorem S_pred: \forall n:nat.lt O n \to n=S (pred n).
+theorem S_pred: \forall n:nat.O<n \to n=S (pred n).
 intro.elim n.apply False_ind.exact not_le_Sn_O O H.
 apply eq_f.apply pred_Sn.
 qed.
 
 (* le vs. lt *)
-theorem lt_to_le : \forall n,m:nat. lt n m \to le n m.
+theorem lt_to_le : \forall n,m:nat. n<m \to n \leq m.
 simplify.intros.elim H.
 apply le_S. apply le_n.
 apply le_S. assumption.
 qed.
 
-theorem not_le_to_lt: \forall n,m:nat. Not (le n m) \to lt m n.
+theorem not_le_to_lt: \forall n,m:nat. \lnot (n \leq m) \to m<n.
 intros 2.
-apply nat_elim2 (\lambda n,m.Not (le n m) \to lt m n).
-intros.apply absurd (le O n1).apply le_O_n.assumption.
+apply nat_elim2 (\lambda n,m.\lnot (n \leq m) \to m<n).
+intros.apply absurd (O \leq n1).apply le_O_n.assumption.
 simplify.intros.apply le_S_S.apply le_O_n.
 simplify.intros.apply le_S_S.apply H.intros.apply H1.apply le_S_S.
 assumption.
 qed.
 
-theorem lt_to_not_le: \forall n,m:nat. lt n m \to Not (le m n).
+theorem lt_to_not_le: \forall n,m:nat. n<m \to \lnot (m \leq n).
 simplify.intros 3.elim H.
 apply not_le_Sn_n n H1.
 apply H2.apply lt_to_le. apply H3.
 qed.
 
 (* le elimination *)
-theorem le_n_O_to_eq : \forall n:nat. (le n O) \to O=n.
+theorem le_n_O_to_eq : \forall n:nat. n \leq O \to O=n.
 intro.elim n.reflexivity.
 apply False_ind.
 (* non si applica not_le_Sn_O *)
 apply  (not_le_Sn_O ? H1).
 qed.
 
-theorem le_n_O_elim: \forall n:nat.le n O \to \forall P: nat \to Prop.
+theorem le_n_O_elim: \forall n:nat.n \leq O \to \forall P: nat \to Prop.
 P O \to P n.
 intro.elim n.
 assumption.
@@ -127,8 +136,8 @@ apply False_ind.
 apply  (not_le_Sn_O ? H1).
 qed.
 
-theorem le_n_Sm_elim : \forall n,m:nat.le n (S m) \to 
-\forall P:Prop. (le (S n) (S m) \to P) \to (n=(S m) \to P) \to P.
+theorem le_n_Sm_elim : \forall n,m:nat.n \leq S m \to 
+\forall P:Prop. (S n \leq S m \to P) \to (n=S m \to P) \to P.
 intros 4.elim H.
 apply H2.reflexivity.
 apply H3. apply le_S_S. assumption.
@@ -137,7 +146,7 @@ qed.
 (* other abstract properties *)
 theorem antisymmetric_le : antisymmetric nat le.
 simplify.intros 2.
-apply nat_elim2 (\lambda n,m.((le n m) \to (le m n) \to n=m)).
+apply nat_elim2 (\lambda n,m.(n \leq m \to m \leq n \to n=m)).
 intros.apply le_n_O_to_eq.assumption.
 intros.apply False_ind.apply not_le_Sn_O ? H.
 intros.apply eq_f.apply H.
@@ -145,12 +154,12 @@ apply le_S_S_to_le.assumption.
 apply le_S_S_to_le.assumption.
 qed.
 
-theorem antisym_le: \forall n,m:nat. le n m \to le m n \to n=m
+theorem antisym_le: \forall n,m:nat. n \leq m \to m \leq n \to n=m
 \def antisymmetric_le.
 
-theorem decidable_le: \forall n,m:nat. decidable (le n m).
+theorem decidable_le: \forall n,m:nat. decidable (n \leq m).
 intros.
-apply nat_elim2 (\lambda n,m.decidable (le n m)).
+apply nat_elim2 (\lambda n,m.decidable (n \leq m)).
 intros.simplify.left.apply le_O_n.
 intros.simplify.right.exact not_le_Sn_O n1.
 intros 2.simplify.intro.elim H.