]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_TPTP/COL053-1.ma
tacticals are really tactics now, they have an AST at the same level of
[helm.git] / helm / software / matita / contribs / ng_TPTP / COL053-1.ma
index 9a4b215d2543c4df4867beeeb341ebfeceb84b73..cdc2673c0f35fadf69bf915f594b5dc876e2e67e 100644 (file)
@@ -68,32 +68,31 @@ include "logic/equality.ma".
 
 (* ----  -[(u)f(u) = A(B((C)f(u)))]. *)
 ntheorem prove_bird_exists:
- ∀Univ:Type.∀U:Univ.∀W:Univ.∀X:Univ.∀Y:Univ.
(∀Univ:Type.∀U:Univ.∀W:Univ.∀X:Univ.∀Y:Univ.
 ∀a:Univ.
 ∀b:Univ.
 ∀c:Univ.
 ∀compose:∀_:Univ.∀_:Univ.Univ.
 ∀f:∀_:Univ.Univ.
 ∀response:∀_:Univ.∀_:Univ.Univ.
-∀H0:∀W:Univ.∀X:Univ.∀Y:Univ.eq Univ (response (compose X Y) W) (response X (response Y W)).∃U:Univ.eq Univ (response U (f U)) (response a (response b (response c (f U))))
+∀H0:∀W:Univ.∀X:Univ.∀Y:Univ.eq Univ (response (compose X Y) W) (response X (response Y W)).∃U:Univ.eq Univ (response U (f U)) (response a (response b (response c (f U)))))
 .
-#Univ.
-#U.
-#W.
-#X.
-#Y.
-#a.
-#b.
-#c.
-#compose.
-#f.
-#response.
-#H0.
-napply ex_intro[
-nid2:
-nauto by H0;
-nid|
-skip]
+#Univ ##.
+#U ##.
+#W ##.
+#X ##.
+#Y ##.
+#a ##.
+#b ##.
+#c ##.
+#compose ##.
+#f ##.
+#response ##.
+#H0 ##.
+napply (ex_intro ? ? ? ?) ##[
+##2:
+nauto by H0 ##;
+##| ##skip ##]
 nqed.
 
 (* -------------------------------------------------------------------------- *)