]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/RELATIONAL/NPlus/props.ma
- new legature == for \equiv used in the notation for NPlus
[helm.git] / helm / software / matita / contribs / RELATIONAL / NPlus / props.ma
index 56146a3bf602d702fe969329930a23f70db6588d..a80e13fa510f316db99284815a34f18794375fb2 100644 (file)
@@ -16,11 +16,12 @@ set "baseuri" "cic:/matita/RELATIONAL/NPlus/props".
 
 include "NPlus/fwd.ma".
 
-theorem nplus_zero_1: \forall q. NPlus zero q q.
+theorem nplus_zero_1: \forall q. zero + q == q.
  intros. elim q; clear q; auto.
 qed.
 
-theorem nplus_succ_1: \forall p,q,r. NPlus p q r \to NPlus (succ p) q (succ r).
+theorem nplus_succ_1: \forall p,q,r. NPlus p q r \to 
+                       (succ p) + q == (succ r).
  intros 2. elim q; clear q;
  [ lapply linear nplus_gen_zero_2 to H as H0.
    rewrite > H0. clear H0 p
@@ -30,7 +31,7 @@ theorem nplus_succ_1: \forall p,q,r. NPlus p q r \to NPlus (succ p) q (succ r).
  ]; auto.
 qed.
 
-theorem nplus_sym: \forall p,q,r. NPlus p q r \to NPlus q p r.
+theorem nplus_sym: \forall p,q,r. (p + q == r) \to q + p == r.
  intros 2. elim q; clear q;
  [ lapply linear nplus_gen_zero_2 to H as H0.
    rewrite > H0. clear H0 p
@@ -41,7 +42,7 @@ theorem nplus_sym: \forall p,q,r. NPlus p q r \to NPlus q p r.
 qed.
 
 theorem nplus_shift_succ_sx: \forall p,q,r. 
-                             NPlus p (succ q) r \to NPlus (succ p) q r.
+                              (p + (succ q) == r) \to (succ p) + q == r.
  intros.
  lapply linear nplus_gen_succ_2 to H as H0.
  decompose.
@@ -50,7 +51,7 @@ theorem nplus_shift_succ_sx: \forall p,q,r.
 qed.
 
 theorem nplus_shift_succ_dx: \forall p,q,r. 
-                             NPlus (succ p) q r \to NPlus p (succ q) r.
+                              ((succ p) + q == r) \to p + (succ q) == r.
  intros.
  lapply linear nplus_gen_succ_1 to H as H0.
  decompose.
@@ -58,9 +59,9 @@ theorem nplus_shift_succ_dx: \forall p,q,r.
  auto.
 qed.
 
-theorem nplus_trans_1: \forall p,q1,r1. NPlus p q1 r1 \to 
-                       \forall q2,r2. NPlus r1 q2 r2 \to
-                       \exists q. NPlus q1 q2 q \land NPlus p q r2.
+theorem nplus_trans_1: \forall p,q1,r1. (p + q1 == r1) \to 
+                        \forall q2,r2. (r1 + q2 == r2) \to
+                        \exists q. (q1 + q2 == q) \land p + q == r2.
  intros 2; elim q1; clear q1; intros;
  [ lapply linear nplus_gen_zero_2 to H as H0.
    rewrite > H0. clear H0 p
@@ -75,9 +76,9 @@ theorem nplus_trans_1: \forall p,q1,r1. NPlus p q1 r1 \to
  ]; apply ex_intro; [| auto || auto ]. (**)
 qed.
 
-theorem nplus_trans_2: \forall p1,q,r1. NPlus p1 q r1 \to 
-                       \forall p2,r2. NPlus p2 r1 r2 \to
-                       \exists p. NPlus p1 p2 p \land NPlus p q r2.
+theorem nplus_trans_2: \forall p1,q,r1. (p1 + q == r1) \to 
+                        \forall p2,r2. (p2 + r1 == r2) \to
+                        \exists p. (p1 + p2 == p) \land p + q == r2.
  intros 2; elim q; clear q; intros;
  [ lapply linear nplus_gen_zero_2 to H as H0.
    rewrite > H0. clear H0 p1
@@ -92,8 +93,8 @@ theorem nplus_trans_2: \forall p1,q,r1. NPlus p1 q r1 \to
  ]; apply ex_intro; [| auto || auto ]. (**)
 qed.
 
-theorem nplus_conf: \forall p,q,r1. NPlus p q r1 \to 
-                    \forall r2. NPlus p q r2 \to r1 = r2.
+theorem nplus_conf: \forall p,q,r1. (p + q == r1) \to 
+                     \forall r2. (p + q == r2) \to r1 = r2.
  intros 2. elim q; clear q; intros;
  [ lapply linear nplus_gen_zero_2 to H as H0.
    rewrite > H0 in H1. clear H0 p