]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_TPTP/COL014-1.ma
made executable again
[helm.git] / helm / software / matita / contribs / ng_TPTP / COL014-1.ma
index a2deaa971fd820367eb576f94f52332c2dc73998..8185750d700d14d123d4140614d5a6321e19c536 100644 (file)
@@ -4,7 +4,7 @@ include "logic/equality.ma".
 
 (* -------------------------------------------------------------------------- *)
 
-(*  File     : COL014-1 : TPTP v3.2.0. Released v1.0.0. *)
+(*  File     : COL014-1 : TPTP v3.7.0. Released v1.0.0. *)
 
 (*  Domain   : Combinatory Logic *)
 
@@ -52,28 +52,28 @@ include "logic/equality.ma".
 
 (* -------------------------------------------------------------------------- *)
 ntheorem prove_fixed_point:
- ∀Univ:Type.∀X:Univ.∀Y:Univ.
(∀Univ:Type.∀X:Univ.∀Y:Univ.
 ∀apply:∀_:Univ.∀_:Univ.Univ.
 ∀combinator:Univ.
 ∀l:Univ.
 ∀o:Univ.
 ∀H0:∀X:Univ.∀Y:Univ.eq Univ (apply (apply o X) Y) (apply Y (apply X Y)).
-∀H1:∀X:Univ.∀Y:Univ.eq Univ (apply (apply l X) Y) (apply X (apply Y Y)).∃Y:Univ.eq Univ Y (apply combinator Y)
+∀H1:∀X:Univ.∀Y:Univ.eq Univ (apply (apply l X) Y) (apply X (apply Y Y)).∃Y:Univ.eq Univ Y (apply combinator Y))
 .
-#Univ.
-#X.
-#Y.
-#apply.
-#combinator.
-#l.
-#o.
-#H0.
-#H1.
-napply ex_intro[
-nid2:
-nauto by H0,H1;
-nid|
-skip]
+#Univ ##.
+#X ##.
+#Y ##.
+#apply ##.
+#combinator ##.
+#l ##.
+#o ##.
+#H0 ##.
+#H1 ##.
+napply (ex_intro ? ? ? ?) ##[
+##2:
+nauto by H0,H1 ##;
+##| ##skip ##]
+ntry (nassumption) ##;
 nqed.
 
 (* -------------------------------------------------------------------------- *)