set "baseuri" "cic:/matita/RELATIONAL-ARITHMETICS/add_props".
-include "nat_props.ma".
-include "add_defs.ma".
+include "add_gen.ma".
-theorem add_gen_O_2: \forall p,r. add p O r \to p = r.
- intros. inversion H; clear H; intros;
- [ reflexivity
- | lapply eq_gen_O_S to H2 as H0. apply H0
- ].
-qed.
-
-theorem add_gen_S_2: \forall p,q,r. add p (S q) r \to
- \exists s. r = (S s) \land add p q s.
- intros. inversion H; clear H; intros;
- [ lapply eq_gen_S_O to H as H0. apply H0
- | lapply eq_gen_S_S to H2 as H0. clear H2.
- rewrite > H0. clear H0.
- apply ex_intro; [| auto ] (**)
- ].
-qed.
-
theorem add_O_1: \forall q. add O q q.
intros. elim q; clear q; auto.
qed.
auto.
qed.
-theorem add_gen_O_1: \forall q,r. add O q r \to q = r.
- intros. auto.
-qed.
-
-theorem add_gen_S_1: \forall p,q,r. add (S p) q r \to
- \exists s. r = (S s) \land add p q s.
- intros. auto.
-qed.
theorem add_shift_S_dx: \forall p,q,r. add (S p) q r \to add p (S q) r.
intros.
rewrite > H2. clear H2. clear r2.
]; auto.
qed.
+
+
+
+theorem add_gen_eq_2_3: \forall p,q. add p q q \to p = O.
+ intros 2. elim q; clear q; intros;
+ [ lapply add_gen_O_2 to H as H0. clear H.
+ rewrite > H0. clear H0. clear p
+ | lapply add_gen_S_2 to H1 as H0. clear H1.
+ decompose H0.
+ lapply eq_gen_S_S to H2 as H0. clear H2.
+ rewrite < H0 in H3. clear H0. clear a
+ ]; auto.
+qed.
+
+theorem add_gen_eq_1_3: \forall p,q. add p q p \to q = O.
+ intros.
+ lapply add_sym to H. clear H.
+ auto.
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+set "baseuri" "cic:/matita/RELATIONAL-ARITHMETICS/nat_gen".
+
+include "library/logic/equality.ma".
+include "nat_defs.ma".
+
+theorem eq_gen_O_S: \forall (P:Prop). \forall m2. O = S m2 \to P.
+ intros. discriminate H.
+qed.
+
+theorem eq_gen_S_O: \forall (P:Prop). \forall m1. S m1 = O \to P.
+ intros. discriminate H.
+qed.
+
+theorem eq_gen_S_S: \forall m1,m2. S m1 = S m2 \to m1 = m2.
+ intros. injection H. assumption.
+qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-set "baseuri" "cic:/matita/RELATIONAL-ARITHMETICS/nat_props".
-
-include "library/logic/equality.ma".
-include "nat_defs.ma".
-
-theorem eq_gen_O_S: \forall (P:Prop). \forall m2. O = S m2 \to P.
- intros. discriminate H.
-qed.
-
-theorem eq_gen_S_O: \forall (P:Prop). \forall m1. S m1 = O \to P.
- intros. discriminate H.
-qed.
-
-theorem eq_gen_S_S: \forall m1,m2. S m1 = S m2 \to m1 = m2.
- intros. injection H. assumption.
-qed.
\ No newline at end of file