]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_TPTP/COL058-1.ma
(no commit message)
[helm.git] / helm / software / matita / contribs / ng_TPTP / COL058-1.ma
index d454217b718c8649e9b435ff08bc7bbe0aefd522..594223ce0097f39ba8d48d0471f09d8d6dc9a744 100644 (file)
@@ -4,7 +4,7 @@ include "logic/equality.ma".
 
 (* -------------------------------------------------------------------------- *)
 
-(*  File     : COL058-1 : TPTP v3.2.0. Released v1.0.0. *)
+(*  File     : COL058-1 : TPTP v3.7.0. Released v1.0.0. *)
 
 (*  Domain   : Combinatory Logic *)
 
@@ -52,23 +52,23 @@ include "logic/equality.ma".
 
 (* ---- Hypothesis: There exists a bird x that is fond of itself.  *)
 ntheorem prove_the_bird_exists:
- ∀Univ:Type.∀X:Univ.∀X1:Univ.∀X2:Univ.
(∀Univ:Type.∀X:Univ.∀X1:Univ.∀X2:Univ.
 ∀lark:Univ.
 ∀response:∀_:Univ.∀_:Univ.Univ.
-∀H0:∀X1:Univ.∀X2:Univ.eq Univ (response (response lark X1) X2) (response X1 (response X2 X2)).∃X:Univ.eq Univ (response X X) X
+∀H0:∀X1:Univ.∀X2:Univ.eq Univ (response (response lark X1) X2) (response X1 (response X2 X2)).∃X:Univ.eq Univ (response X X) X)
 .
-#Univ.
-#X.
-#X1.
-#X2.
-#lark.
-#response.
-#H0.
-napply ex_intro[
-nid2:
-nauto by H0;
-nid|
-skip]
+#Univ ##.
+#X ##.
+#X1 ##.
+#X2 ##.
+#lark ##.
+#response ##.
+#H0 ##.
+napply (ex_intro ? ? ? ?) ##[
+##2:
+nauto by H0 ##;
+##| ##skip ##]
+ntry (nassumption) ##;
 nqed.
 
 (* -------------------------------------------------------------------------- *)