]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/tests/unfold.ma
ocaml 3.09 transition
[helm.git] / helm / matita / tests / unfold.ma
index d63dd50d46d8a6a8027234ff691815f4a71c537b..d17e5a2daf9fd243c136f10cfcd6997707226ede 100644 (file)
@@ -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,6 +34,7 @@ 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.