]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/RELATIONAL/NPlus/fwd.ma
Added timeout to autos here and there.
[helm.git] / helm / software / matita / contribs / RELATIONAL / NPlus / fwd.ma
index bae48d4051d774dccb0812918489f53243247139..0b8ba2b889b35456329db195600137dee4dbed4d 100644 (file)
@@ -19,69 +19,69 @@ include "NPlus/defs.ma".
 
 (* primitive generation lemmas proved by elimination and inversion *)
 
-theorem nplus_gen_zero_1: \forall q,r. NPlus zero q r \to q = r.
+theorem nplus_gen_zero_1: \forall q,r. (zero + q == r) \to q = r.
  intros. elim H; clear H q r; intros;
  [ reflexivity
- | clear H1. auto
+ | clear H1. auto new timeout=30
  ].
 qed.
 
-theorem nplus_gen_succ_1: \forall p,q,r. NPlus (succ p) q r \to 
-                          \exists s. r = (succ s) \land NPlus p q s.
+theorem nplus_gen_succ_1: \forall p,q,r. ((succ p) + q == r) \to 
+                          \exists s. r = (succ s) \land p + q == s.
  intros. elim H; clear H q r; intros;
  [
  | clear H1.
    decompose.
-   rewrite > H1. clear H1 n2
- ]; apply ex_intro; [| auto || auto ]. (**)
+   subst.
+ ]; apply ex_intro; [| auto new timeout=30 || auto new timeout=30 ]. (**)
 qed.
 
-theorem nplus_gen_zero_2: \forall p,r. NPlus p zero r \to p = r.
+theorem nplus_gen_zero_2: \forall p,r. (p + zero == r) \to p = r.
  intros. inversion H; clear H; intros;
- [ auto
+ [ auto new timeout=30
  | clear H H1.
    lapply eq_gen_zero_succ to H2 as H0. apply H0
  ].
 qed.
 
-theorem nplus_gen_succ_2: \forall p,q,r. NPlus p (succ q) r \to 
-                          \exists s. r = (succ s) \land NPlus p q s.
+theorem nplus_gen_succ_2: \forall p,q,r. (p + (succ q) == r) \to 
+                          \exists s. r = (succ s) \land p + q == s.
  intros. inversion H; clear H; intros;
  [ lapply eq_gen_succ_zero to H as H0. apply H0
  | clear H1 H3 r.
    lapply linear eq_gen_succ_succ to H2 as H0.
-   rewrite > H0. clear H0 q.
-   apply ex_intro; [| auto ] (**)
+   subst.
+   apply ex_intro; [| auto new timeout=30 ] (**)
  ].
 qed.
 
-theorem nplus_gen_zero_3: \forall p,q. NPlus p q zero \to p = zero \land q = zero.
+theorem nplus_gen_zero_3: \forall p,q. (p + q == zero) \to 
+                          p = zero \land q = zero.
  intros. inversion H; clear H; intros;
- [ rewrite < H1. clear H1 p.
-   auto
+ [ subst. auto new timeout=30
  | clear H H1.
    lapply eq_gen_zero_succ to H3 as H0. apply H0
  ].
 qed.
 
-theorem nplus_gen_succ_3: \forall p,q,r. NPlus p q (succ r) \to
-                          \exists s. p = succ s \land NPlus s q r \lor
-                                     q = succ s \land NPlus p s r.
+theorem nplus_gen_succ_3: \forall p,q,r. (p + q == (succ r)) \to
+                          \exists s. p = succ s \land (s + q == r) \lor
+                                     q = succ s \land p + s == r.
  intros. inversion H; clear H; intros;
- [ rewrite < H1. clear H1 p
+ [ subst.
  | clear H1.
    lapply linear eq_gen_succ_succ to H3 as H0.
-   rewrite > H0. clear H0 r.
- ]; apply ex_intro; [| auto || auto ] (**)
+   subst.
+ ]; apply ex_intro; [| auto new timeout=30 || auto new timeout=30 ] (**)
 qed.
 (*
 (* alternative proofs invoking nplus_gen_2 *)
 
-variant nplus_gen_zero_3_alt: \forall p,q. NPlus p q zero \to p = zero \land q = zero.
+variant nplus_gen_zero_3_alt: \forall p,q. (p + q == zero) \to 
+                              p = zero \land q = zero.
  intros 2. elim q; clear q; intros;
  [ lapply linear nplus_gen_zero_2 to H as H0.
-   rewrite > H0. clear H0 p.
-   auto
+   subst. auto new timeout=30
  | clear H.
    lapply linear nplus_gen_succ_2 to H1 as H0.
    decompose.
@@ -89,40 +89,40 @@ variant nplus_gen_zero_3_alt: \forall p,q. NPlus p q zero \to p = zero \land q =
  ].
 qed.
 
-variant nplus_gen_succ_3_alt: \forall p,q,r. NPlus p q (succ r) \to
-                              \exists s. p = succ s \land NPlus s q r \lor
-                                         q = succ s \land NPlus p s r.
+variant nplus_gen_succ_3_alt: \forall p,q,r. (p + q == (succ r)) \to
+                              \exists s. p = succ s \land (s + q == r) \lor
+                                         q = succ s \land p + s == r.
  intros 2. elim q; clear q; intros;
  [ lapply linear nplus_gen_zero_2 to H as H0.
-   rewrite > H0. clear H0 p
+   subst
  | clear H.
    lapply linear nplus_gen_succ_2 to H1 as H0.
    decompose.
    lapply linear eq_gen_succ_succ to H1 as H0.
-   rewrite > H0. clear H0 r.
- ]; apply ex_intro; [| auto || auto ]. (**)
+   subst
+ ]; apply ex_intro; [| auto new timeout=30 || auto new timeout=30 ]. (**)
 qed.
 *)
 (* other simplification lemmas *)
 
-theorem nplus_gen_eq_2_3: \forall p,q. NPlus p q q \to p = zero.
+theorem nplus_gen_eq_2_3: \forall p,q. (p + q == q) \to p = zero.
  intros 2. elim q; clear q; intros;
  [ lapply linear nplus_gen_zero_2 to H as H0.
-   rewrite > H0. clear H0 p
+   subst
  | lapply linear nplus_gen_succ_2 to H1 as H0.
    decompose.
    lapply linear eq_gen_succ_succ to H2 as H0.
-   rewrite < H0 in H3. clear H0 a
- ]; auto.
+   subst
+ ]; auto new timeout=30.
 qed.
 
-theorem nplus_gen_eq_1_3: \forall p,q. NPlus p q p \to q = zero.
+theorem nplus_gen_eq_1_3: \forall p,q. (p + q == p) \to q = zero.
  intros 1. elim p; clear p; intros;
  [ lapply linear nplus_gen_zero_1 to H as H0.
-   rewrite > H0. clear H0 q
+   subst
  | lapply linear nplus_gen_succ_1 to H1 as H0.
    decompose.
    lapply linear eq_gen_succ_succ to H2 as H0.
-   rewrite < H0 in H3. clear H0 a
- ]; auto.
+   subst
+ ]; auto new timeout=30.
 qed.