]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_TPTP/SYN305-1.ma
tacticals are really tactics now, they have an AST at the same level of
[helm.git] / helm / software / matita / contribs / ng_TPTP / SYN305-1.ma
index 938dd75b77999a008440e5553cf01a9867f7793c..bd361c182144efc78b5abdbe9f4d85d50fc529d2 100644 (file)
@@ -42,25 +42,24 @@ include "logic/equality.ma".
 
 (* -------------------------------------------------------------------------- *)
 ntheorem clause3:
- ∀Univ:Type.∀X:Univ.
(∀Univ:Type.∀X:Univ.
 ∀f:∀_:Univ.Univ.
 ∀g1:∀_:Univ.Univ.
 ∀g2:∀_:Univ.Univ.
 ∀H0:∀X:Univ.eq Univ (f (g2 X)) X.
-∀H1:∀X:Univ.eq Univ (f (g1 X)) X.∃X:Univ.eq Univ (g1 X) (g2 X)
+∀H1:∀X:Univ.eq Univ (f (g1 X)) X.∃X:Univ.eq Univ (g1 X) (g2 X))
 .
-#Univ.
-#X.
-#f.
-#g1.
-#g2.
-#H0.
-#H1.
-napply ex_intro[
-nid2:
-nauto by H0,H1;
-nid|
-skip]
+#Univ ##.
+#X ##.
+#f ##.
+#g1 ##.
+#g2 ##.
+#H0 ##.
+#H1 ##.
+napply (ex_intro ? ? ? ?) ##[
+##2:
+nauto by H0,H1 ##;
+##| ##skip ##]
 nqed.
 
 (* -------------------------------------------------------------------------- *)