X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Ftests%2Funfold.ma;h=d17e5a2daf9fd243c136f10cfcd6997707226ede;hb=4167cea65ca58897d1a3dbb81ff95de5074700cc;hp=4af22f1b97c03796db9dc34dc34fd94099bc16b5;hpb=0c2b65693a5a7a2881ebb6dfcaa432f4f9fd22a4;p=helm.git diff --git a/helm/matita/tests/unfold.ma b/helm/matita/tests/unfold.ma index 4af22f1b9..d17e5a2da 100644 --- a/helm/matita/tests/unfold.ma +++ b/helm/matita/tests/unfold.ma @@ -19,12 +19,13 @@ include "coq.ma". alias symbol "plus" (instance 0) = "Coq's natural plus". definition myplus \def \lambda x,y. x+y. +alias id "S" = "cic:/Coq/Init/Datatypes/nat.ind#xpointer(1/1/2)". lemma lem: \forall n. S (n + n) = (S n) + n. intro; reflexivity. qed. theorem trivial: \forall n. S (myplus n n) = myplus (S n) n. - unfold myplus in \vdash \forall _.(? ? ? %). + unfold myplus in \vdash (\forall _.(? ? ? %)). intro. unfold myplus. rewrite > lem. @@ -33,5 +34,8 @@ qed. (* This test needs to parse "uno" in the context of the hypothesis H, not in the context of the goal. *) +alias id "O" = "cic:/Coq/Init/Datatypes/nat.ind#xpointer(1/1/1)". theorem t: let uno \def S O in uno + uno = S uno \to uno=uno. intros. unfold uno in H. + reflexivity. +qed.