]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_TPTP/COL052-1.ma
tacticals are really tactics now, they have an AST at the same level of
[helm.git] / helm / software / matita / contribs / ng_TPTP / COL052-1.ma
index c207b57dc1e154d1c7d33057f0bf06faa661d30e..0a3889b2d9f93b5d017b6a885e3df8d36272c4da 100644 (file)
@@ -74,7 +74,7 @@ include "logic/equality.ma".
 
 (* ----      -(response(A,v) = response(odd_bird,v)). *)
 ntheorem prove_a_is_agreeable:
- ∀Univ:Type.∀V:Univ.∀W:Univ.∀X:Univ.∀Y:Univ.
(∀Univ:Type.∀V:Univ.∀W:Univ.∀X:Univ.∀Y:Univ.
 ∀a:Univ.
 ∀c:Univ.
 ∀common_bird:∀_:Univ.Univ.
@@ -82,26 +82,25 @@ ntheorem prove_a_is_agreeable:
 ∀odd_bird:Univ.
 ∀response:∀_:Univ.∀_:Univ.Univ.
 ∀H0:∀X:Univ.eq Univ (response c (common_bird X)) (response X (common_bird X)).
-∀H1:∀W:Univ.∀X:Univ.∀Y:Univ.eq Univ (response (compose X Y) W) (response X (response Y W)).∃V:Univ.eq Univ (response a V) (response odd_bird V)
+∀H1:∀W:Univ.∀X:Univ.∀Y:Univ.eq Univ (response (compose X Y) W) (response X (response Y W)).∃V:Univ.eq Univ (response a V) (response odd_bird V))
 .
-#Univ.
-#V.
-#W.
-#X.
-#Y.
-#a.
-#c.
-#common_bird.
-#compose.
-#odd_bird.
-#response.
-#H0.
-#H1.
-napply ex_intro[
-nid2:
-nauto by H0,H1;
-nid|
-skip]
+#Univ ##.
+#V ##.
+#W ##.
+#X ##.
+#Y ##.
+#a ##.
+#c ##.
+#common_bird ##.
+#compose ##.
+#odd_bird ##.
+#response ##.
+#H0 ##.
+#H1 ##.
+napply (ex_intro ? ? ? ?) ##[
+##2:
+nauto by H0,H1 ##;
+##| ##skip ##]
 nqed.
 
 (* ----C composes A with B. WHY is this here?  *)