]> matita.cs.unibo.it Git - helm.git/commitdiff
auto new => auto new library in those tests that use Coq's library
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Mon, 27 Nov 2006 10:44:18 +0000 (10:44 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Mon, 27 Nov 2006 10:44:18 +0000 (10:44 +0000)
matita/tests/elim.ma
matita/tests/inversion.ma
matita/tests/replace.ma
matita/tests/test3.ma

index 4681d83597a0de52fde58a85dcde4da2b859e372..d6a4c3f33f7e994f362100e40ec391b0702ce569 100644 (file)
@@ -76,5 +76,5 @@ theorem t': \forall x,y. \forall H: sum x y O.
  simplify. intros.
  generalize in match H1.
  rewrite < H2; rewrite < H3.intro.
- rewrite > H4.auto new.
+ rewrite > H4.auto new library.
 qed.
index 35963ab7392bf2b99fba801b6fb3ebf54795fc42..925c54677ed4d0f61fd2b242caa9c9e09025e39e 100644 (file)
@@ -42,7 +42,7 @@ theorem t: \forall x,y. \forall H: sum x y O.
  simplify. intros.
  generalize in match H1.
  rewrite < H2; rewrite < H3.intro.
- rewrite > H4.auto new.
+ rewrite > H4.auto new library.
 qed.
 
 theorem t1: \forall x,y. sum x y O \to x = y.
@@ -57,5 +57,5 @@ apply (sum_ind ? (\lambda a,b,K. y=a \to O=b \to x=a) ? ? ? s).*)
 inversion s.
 intros.simplify.
 intros.
-rewrite > H. rewrite < H2.  auto new.
+rewrite > H. rewrite < H2.  auto new library.
 qed.
index e9960a507de28d4ee7e3dcb344e2b2e373df12d4..b740f387ef032e5f6c795ecdac86bcffe10d7f9e 100644 (file)
@@ -30,7 +30,7 @@ theorem t: \forall x:nat. x * (x + 0) = (0 + x) * (x + x * 0).
  rewrite < (plus_n_O x).
  reflexivity.
  reflexivity.
- auto new.
+ auto new library.
 qed.
 
 (* This test tests "replace in match t" where t contains some metavariables *)
index 1399f8be20e745f5d45629f694392c7b7da8f6af..08d4cbcb6a0f49783da0959e0af7b1779063391e 100644 (file)
@@ -27,5 +27,5 @@ alias symbol "times" (instance 0) = "Coq's natural times".
 
 theorem b:\forall p:nat. p * 0=0.
 intro.
-auto new.
+auto new library.
 qed.