]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_TPTP/COL085-1.ma
tacticals are really tactics now, they have an AST at the same level of
[helm.git] / helm / software / matita / contribs / ng_TPTP / COL085-1.ma
index c42c180089ad79aad03afe63410283db905976a0..a1aba169f57b5a6adfeedee67b124a13669436e0 100644 (file)
@@ -42,28 +42,26 @@ include "logic/equality.ma".
 
 (* -------------------------------------------------------------------------- *)
 ntheorem prove_happiness_1:
- ∀Univ:Type.∀A:Univ.∀B:Univ.
(∀Univ:Type.∀A:Univ.∀B:Univ.
 ∀a:Univ.
 ∀b:Univ.
 ∀response:∀_:Univ.∀_:Univ.Univ.
-∀H0:eq Univ (response a b) b.∃A:Univ.∃B:Univ.eq Univ (response a A) B
+∀H0:eq Univ (response a b) b.∃A:Univ.∃B:Univ.eq Univ (response a A) B)
 .
-#Univ.
-#A.
-#B.
-#a.
-#b.
-#response.
-#H0.
-napply ex_intro[
-nid2:
-napply ex_intro[
-nid2:
-nauto by H0;
-nid|
-skip]
-nid|
-skip]
+#Univ ##.
+#A ##.
+#B ##.
+#a ##.
+#b ##.
+#response ##.
+#H0 ##.
+napply (ex_intro ? ? ? ?) ##[
+##2:
+napply (ex_intro ? ? ? ?) ##[
+##2:
+nauto by H0 ##;
+##| ##skip ##]
+##| ##skip ##]
 nqed.
 
 (* -------------------------------------------------------------------------- *)