]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_TPTP/COL056-1.ma
tacticals are really tactics now, they have an AST at the same level of
[helm.git] / helm / software / matita / contribs / ng_TPTP / COL056-1.ma
index fc787c7db3d1fa5b025ea7bf9df5123d9245957c..c12e91f31130da65e61ba143b8b577971bc24880 100644 (file)
@@ -70,7 +70,7 @@ include "logic/equality.ma".
 
 (* ----   (AB = C) and (AC = B) and -(wv = v). *)
 ntheorem prove_there_exists_a_happy_bird:
- ∀Univ:Type.∀V:Univ.∀W:Univ.∀X:Univ.∀Y:Univ.
(∀Univ:Type.∀V:Univ.∀W:Univ.∀X:Univ.∀Y:Univ.
 ∀a:Univ.
 ∀b:Univ.
 ∀c:Univ.
@@ -78,30 +78,28 @@ ntheorem prove_there_exists_a_happy_bird:
 ∀response:∀_:Univ.∀_:Univ.Univ.
 ∀H0:eq Univ (response a c) b.
 ∀H1:eq Univ (response a b) c.
-∀H2:∀W:Univ.∀X:Univ.∀Y:Univ.eq Univ (response (compose X Y) W) (response X (response Y W)).∃V:Univ.∃W:Univ.eq Univ (response W V) V
+∀H2:∀W:Univ.∀X:Univ.∀Y:Univ.eq Univ (response (compose X Y) W) (response X (response Y W)).∃V:Univ.∃W:Univ.eq Univ (response W V) V)
 .
-#Univ.
-#V.
-#W.
-#X.
-#Y.
-#a.
-#b.
-#c.
-#compose.
-#response.
-#H0.
-#H1.
-#H2.
-napply ex_intro[
-nid2:
-napply ex_intro[
-nid2:
-nauto by H0,H1,H2;
-nid|
-skip]
-nid|
-skip]
+#Univ ##.
+#V ##.
+#W ##.
+#X ##.
+#Y ##.
+#a ##.
+#b ##.
+#c ##.
+#compose ##.
+#response ##.
+#H0 ##.
+#H1 ##.
+#H2 ##.
+napply (ex_intro ? ? ? ?) ##[
+##2:
+napply (ex_intro ? ? ? ?) ##[
+##2:
+nauto by H0,H1,H2 ##;
+##| ##skip ##]
+##| ##skip ##]
 nqed.
 
 (* -------------------------------------------------------------------------- *)