]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma
notational update in ground_2 and basic_2
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / lib / arith.ma
index e3e22484394c91625310912317448ae62121db81..9135ff24d72787e8835f9d1cb7b75898f4c4f2d0 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "ground_2/notation/constructors/successor_1.ma".
-include "ground_2/notation/functions/predecessor_1.ma".
+include "ground_2/notation/constructors/uparrow_1.ma".
+include "ground_2/notation/functions/downarrow_1.ma".
 include "arithmetics/nat.ma".
 include "ground_2/lib/relations.ma".
 
 (* ARITHMETICAL PROPERTIES **************************************************)
 
-interpretation "nat successor" 'Successor m = (S m).
+interpretation "nat successor" 'UpArrow m = (S m).
 
-interpretation "nat predecessor" 'Predecessor m = (pred m).
+interpretation "nat predecessor" 'DownArrow m = (pred m).
 
 interpretation "nat min" 'and x y = (min x y).
 
@@ -35,7 +35,7 @@ normalize // qed.
 lemma pred_S: ∀m. pred (S m) = m.
 // qed.
 
-lemma plus_S1: â\88\80x,y. â«¯(x+y) = (⫯x) + y.
+lemma plus_S1: â\88\80x,y. â\86\91(x+y) = (â\86\91x) + y.
 // qed.
 
 lemma max_O1: ∀n. n = (0 ∨ n).
@@ -44,14 +44,14 @@ lemma max_O1: ∀n. n = (0 ∨ n).
 lemma max_O2: ∀n. n = (n ∨ 0).
 // qed.
 
-lemma max_SS: â\88\80n1,n2. â«¯(n1â\88¨n2) = (⫯n1 â\88¨ â«¯n2).
+lemma max_SS: â\88\80n1,n2. â\86\91(n1â\88¨n2) = (â\86\91n1 â\88¨ â\86\91n2).
 #n1 #n2 elim (decidable_le n1 n2) #H normalize
 [ >(le_to_leb_true … H) | >(not_le_to_leb_false … H) ] -H //
 qed.
 
 (* Equations ****************************************************************)
 
-lemma plus_SO: â\88\80n. n + 1 = â«¯n.
+lemma plus_SO: â\88\80n. n + 1 = â\86\91n.
 // qed.
 
 lemma minus_plus_m_m_commutative: ∀n,m:nat. n = m + n - m.
@@ -145,16 +145,16 @@ lemma monotonic_lt_pred: ∀m,n. m < n → 0 < m → pred m < pred n.
 @le_S_S_to_le >S_pred /2 width=3 by transitive_lt/
 qed.
 
-lemma lt_S_S: â\88\80x,y. x < y â\86\92 â«¯x < â«¯y.
+lemma lt_S_S: â\88\80x,y. x < y â\86\92 â\86\91x < â\86\91y.
 /2 width=1 by le_S_S/ qed.
 
-lemma lt_S: â\88\80n,m. n < m â\86\92 n < â«¯m.
+lemma lt_S: â\88\80n,m. n < m â\86\92 n < â\86\91m.
 /2 width=1 by le_S/ qed.
 
-lemma max_S1_le_S: â\88\80n1,n2,n. (n1 â\88¨ n2) â\89¤ n â\86\92 (⫯n1 â\88¨ n2) â\89¤ â«¯n.
+lemma max_S1_le_S: â\88\80n1,n2,n. (n1 â\88¨ n2) â\89¤ n â\86\92 (â\86\91n1 â\88¨ n2) â\89¤ â\86\91n.
 /4 width=2 by to_max, le_maxr, le_S_S, le_S/ qed-.
 
-lemma max_S2_le_S: â\88\80n1,n2,n. (n1 â\88¨ n2) â\89¤ n â\86\92 (n1 â\88¨ â«¯n2) â\89¤ â«¯n.
+lemma max_S2_le_S: â\88\80n1,n2,n. (n1 â\88¨ n2) â\89¤ n â\86\92 (n1 â\88¨ â\86\91n2) â\89¤ â\86\91n.
 /2 width=1 by max_S1_le_S/ qed-.
 
 lemma arith_j: ∀x,y,z. x-y-1 ≤ x-(y-z)-1.
@@ -187,8 +187,8 @@ qed-.
 lemma lt_le_false: ∀x,y. x < y → y ≤ x → ⊥.
 /3 width=4 by lt_refl_false, lt_to_le_to_lt/ qed-.
 
-lemma succ_inv_refl_sn: â\88\80x. â«¯x = x → ⊥.
-#x #H @(lt_le_false x (⫯x)) //
+lemma succ_inv_refl_sn: â\88\80x. â\86\91x = x → ⊥.
+#x #H @(lt_le_false x (â\86\91x)) //
 qed-.
 
 lemma le_plus_xSy_O_false: ∀x,y. x + S y ≤ 0 → ⊥.
@@ -225,7 +225,7 @@ lapply (monotonic_le_plus_l n1 … Hm) -Hm >H -H
 /2 width=2 by le_plus_to_le/
 qed-.
 
-lemma lt_S_S_to_lt: â\88\80x,y. â«¯x < â«¯y → x < y.
+lemma lt_S_S_to_lt: â\88\80x,y. â\86\91x < â\86\91y → x < y.
 /2 width=1 by le_S_S_to_le/ qed-.
 
 (* Note this should go in nat.ma *)
@@ -237,17 +237,17 @@ lemma discr_x_minus_xy: ∀x,y. x = x - y → x = 0 ∨ y = 0.
 #H destruct
 qed-.
 
-lemma lt_inv_O1: â\88\80n. 0 < n â\86\92 â\88\83m. â«¯m = n.
+lemma lt_inv_O1: â\88\80n. 0 < n â\86\92 â\88\83m. â\86\91m = n.
 * /2 width=2 by ex_intro/
 #H cases (lt_le_false … H) -H //
 qed-.
 
-lemma lt_inv_S1: â\88\80m,n. â«¯m < n â\86\92 â\88\83â\88\83p. m < p & â«¯p = n.
+lemma lt_inv_S1: â\88\80m,n. â\86\91m < n â\86\92 â\88\83â\88\83p. m < p & â\86\91p = n.
 #m * /3 width=3 by lt_S_S_to_lt, ex2_intro/
 #H cases (lt_le_false … H) -H //
 qed-.
 
-lemma lt_inv_gen: â\88\80y,x. x < y â\86\92 â\88\83â\88\83z. x â\89¤ z & â«¯z = y.
+lemma lt_inv_gen: â\88\80y,x. x < y â\86\92 â\88\83â\88\83z. x â\89¤ z & â\86\91z = y.
 * /3 width=3 by le_S_S_to_le, ex2_intro/
 #x #H elim (lt_le_false … H) -H //
 qed-.
@@ -263,13 +263,13 @@ lemma zero_eq_plus: ∀x,y. 0 = x + y → 0 = x ∧ 0 = y.
 * /2 width=1 by conj/ #x #y normalize #H destruct
 qed-.
 
-lemma nat_split: â\88\80x. x = 0 â\88¨ â\88\83y. â«¯y = x.
+lemma nat_split: â\88\80x. x = 0 â\88¨ â\88\83y. â\86\91y = x.
 * /3 width=2 by ex_intro, or_introl, or_intror/
 qed-.
 
 lemma lt_elim: ∀R:relation nat.
-               (â\88\80n2. R O (⫯n2)) →
-               (â\88\80n1,n2. R n1 n2 â\86\92 R (⫯n1) (⫯n2)) →
+               (â\88\80n2. R O (â\86\91n2)) →
+               (â\88\80n1,n2. R n1 n2 â\86\92 R (â\86\91n1) (â\86\91n2)) →
                ∀n2,n1. n1 < n2 → R n1 n2.
 #R #IH1 #IH2 #n2 elim n2 -n2
 [ #n1 #H elim (lt_le_false … H) -H //
@@ -279,7 +279,7 @@ qed-.
 
 lemma le_elim: ∀R:relation nat.
                (∀n2. R O (n2)) →
-               (â\88\80n1,n2. R n1 n2 â\86\92 R (⫯n1) (⫯n2)) →
+               (â\88\80n1,n2. R n1 n2 â\86\92 R (â\86\91n1) (â\86\91n2)) →
                ∀n1,n2. n1 ≤ n2 → R n1 n2.
 #R #IH1 #IH2 #n1 #n2 @(nat_elim2 … n1 n2) -n1 -n2
 /4 width=1 by monotonic_pred/ -IH1 -IH2