X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Ftests%2Finversion.ma;h=c76851f59fa8dd1240d87fdb04d8072a5544ba5f;hb=12cc5b2b8e7f7bb0b5e315094b008a293a4df6b1;hp=8f57534be4d0ffa0159bc86fea7f7abddd2ecaa7;hpb=2c2fd391de336d5611f1739993c31cfb42e7335e;p=helm.git diff --git a/helm/matita/tests/inversion.ma b/helm/matita/tests/inversion.ma index 8f57534be..c76851f59 100644 --- a/helm/matita/tests/inversion.ma +++ b/helm/matita/tests/inversion.ma @@ -1,3 +1,5 @@ +set "baseuri" "cic:/matita/tests/inversion/". + inductive nat : Set \def O : nat | S : nat \to nat. @@ -15,10 +17,10 @@ theorem test_inversion: \forall n. le n O \to n=O. (* goal 2: 0 = 0 *) goal 7. reflexivity. (* goal 1 *) - generalize Hcut. + generalize in match Hcut. apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H). (* first goal (left open) *) - intro. rewrite right H1. + intro. rewrite < H1. clear Hcut. (* second goal (closed) *) goal 14. @@ -33,10 +35,10 @@ alias id "refl_equal" = "cic:/Coq/Init/Logic/eq.ind#xpointer(1/1/1)". theorem test_inversion2: \forall n. le n O \to n=O. intros. (* inversion begins *) - generalize (refl_equal nat O). + generalize in match (refl_equal nat O). apply (le_ind ? (\lambda x. O=x \to n=x) ? ? ? H). (* first goal (left open) *) - intro. rewrite right H1. + intro. rewrite < H1. (* second goal (closed) *) goal 13. simplify. intros.