]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_TPTP/COL032-1.ma
(no commit message)
[helm.git] / helm / software / matita / contribs / ng_TPTP / COL032-1.ma
index ee19a662c1272a755e94ed4cc4cc0785da296d71..7ba259fbf311f091789c990f0f64d48b7ad70bda 100644 (file)
@@ -4,7 +4,7 @@ include "logic/equality.ma".
 
 (* -------------------------------------------------------------------------- *)
 
-(*  File     : COL032-1 : TPTP v3.2.0. Released v1.0.0. *)
+(*  File     : COL032-1 : TPTP v3.7.0. Released v1.0.0. *)
 
 (*  Domain   : Combinatory Logic *)
 
@@ -32,7 +32,7 @@ include "logic/equality.ma".
 
 (*  Status   : Unsatisfiable *)
 
-(*  Rating   : 0.14 v3.1.0, 0.22 v2.7.0, 0.00 v2.2.1, 0.11 v2.2.0, 0.00 v2.1.0, 0.13 v2.0.0 *)
+(*  Rating   : 0.11 v3.4.0, 0.12 v3.3.0, 0.14 v3.1.0, 0.22 v2.7.0, 0.00 v2.2.1, 0.11 v2.2.0, 0.00 v2.1.0, 0.13 v2.0.0 *)
 
 (*  Syntax   : Number of clauses     :    3 (   0 non-Horn;   3 unit;   1 RR) *)
 
@@ -52,29 +52,29 @@ include "logic/equality.ma".
 
 (* -------------------------------------------------------------------------- *)
 ntheorem prove_fixed_point:
- ∀Univ:Type.∀X:Univ.∀Y:Univ.∀Z:Univ.
(∀Univ:Type.∀X:Univ.∀Y:Univ.∀Z:Univ.
 ∀apply:∀_:Univ.∀_:Univ.Univ.
 ∀f:∀_:Univ.Univ.
 ∀m:Univ.
 ∀q:Univ.
 ∀H0:∀X:Univ.∀Y:Univ.∀Z:Univ.eq Univ (apply (apply (apply q X) Y) Z) (apply Y (apply X Z)).
-∀H1:∀X:Univ.eq Univ (apply m X) (apply X X).∃Y:Univ.eq Univ (apply Y (f Y)) (apply (f Y) (apply Y (f Y)))
+∀H1:∀X:Univ.eq Univ (apply m X) (apply X X).∃Y:Univ.eq Univ (apply Y (f Y)) (apply (f Y) (apply Y (f Y))))
 .
-#Univ.
-#X.
-#Y.
-#Z.
-#apply.
-#f.
-#m.
-#q.
-#H0.
-#H1.
-napply ex_intro[
-nid2:
-nauto by H0,H1;
-nid|
-skip]
+#Univ ##.
+#X ##.
+#Y ##.
+#Z ##.
+#apply ##.
+#f ##.
+#m ##.
+#q ##.
+#H0 ##.
+#H1 ##.
+napply (ex_intro ? ? ? ?) ##[
+##2:
+nauto by H0,H1 ##;
+##| ##skip ##]
+ntry (nassumption) ##;
 nqed.
 
 (* -------------------------------------------------------------------------- *)