]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_TPTP/COL049-1.ma
Regenerated problems with corrected tptp2grafite
[helm.git] / helm / software / matita / contribs / ng_TPTP / COL049-1.ma
index 169eaaa40ab20a0b2e86087a5532f20d2c3ca266..120d93770113a6e3df5532511b78d22dbd2a849b 100644 (file)
@@ -4,7 +4,7 @@ include "logic/equality.ma".
 
 (* -------------------------------------------------------------------------- *)
 
-(*  File     : COL049-1 : TPTP v3.2.0. Released v1.0.0. *)
+(*  File     : COL049-1 : TPTP v3.7.0. Released v1.0.0. *)
 
 (*  Domain   : Combinatory Logic *)
 
@@ -52,7 +52,7 @@ include "logic/equality.ma".
 
 (*  Status   : Unsatisfiable *)
 
-(*  Rating   : 0.29 v3.1.0, 0.44 v2.7.0, 0.27 v2.6.0, 0.17 v2.5.0, 0.25 v2.4.0, 0.00 v2.2.1, 0.22 v2.2.0, 0.14 v2.1.0, 0.62 v2.0.0 *)
+(*  Rating   : 0.22 v3.4.0, 0.25 v3.3.0, 0.29 v3.1.0, 0.44 v2.7.0, 0.27 v2.6.0, 0.17 v2.5.0, 0.25 v2.4.0, 0.00 v2.2.1, 0.22 v2.2.0, 0.14 v2.1.0, 0.62 v2.0.0 *)
 
 (*  Syntax   : Number of clauses     :    4 (   0 non-Horn;   4 unit;   1 RR) *)
 
@@ -72,7 +72,7 @@ include "logic/equality.ma".
 
 (* -------------------------------------------------------------------------- *)
 ntheorem prove_strong_fixed_point:
- ∀Univ:Type.∀X:Univ.∀Y:Univ.∀Z:Univ.
(∀Univ:Type.∀X:Univ.∀Y:Univ.∀Z:Univ.
 ∀apply:∀_:Univ.∀_:Univ.Univ.
 ∀b:Univ.
 ∀f:∀_:Univ.Univ.
@@ -80,25 +80,25 @@ ntheorem prove_strong_fixed_point:
 ∀w:Univ.
 ∀H0:∀X:Univ.eq Univ (apply m X) (apply X X).
 ∀H1:∀X:Univ.∀Y:Univ.eq Univ (apply (apply w X) Y) (apply (apply X Y) Y).
-∀H2:∀X:Univ.∀Y:Univ.∀Z:Univ.eq Univ (apply (apply (apply b X) Y) Z) (apply X (apply Y Z)).∃Y:Univ.eq Univ (apply Y (f Y)) (apply (f Y) (apply Y (f Y)))
+∀H2:∀X:Univ.∀Y:Univ.∀Z:Univ.eq Univ (apply (apply (apply b X) Y) Z) (apply X (apply Y Z)).∃Y:Univ.eq Univ (apply Y (f Y)) (apply (f Y) (apply Y (f Y))))
 .
-#Univ.
-#X.
-#Y.
-#Z.
-#apply.
-#b.
-#f.
-#m.
-#w.
-#H0.
-#H1.
-#H2.
-napply ex_intro[
-nid2:
-nauto by H0,H1,H2;
-nid|
-skip]
+#Univ ##.
+#X ##.
+#Y ##.
+#Z ##.
+#apply ##.
+#b ##.
+#f ##.
+#m ##.
+#w ##.
+#H0 ##.
+#H1 ##.
+#H2 ##.
+napply (ex_intro ? ? ? ?) ##[
+##2:
+nauto by H0,H1,H2 ##;
+##| ##skip ##]
+ntry (nassumption) ##;
 nqed.
 
 (* -------------------------------------------------------------------------- *)