1 include "logic/equality.ma".
3 (* Inclusion of: LDA007-3.p *)
5 (* -------------------------------------------------------------------------- *)
7 (* File : LDA007-3 : TPTP v3.2.0. Released v1.0.0. *)
9 (* Domain : LD-Algebras (Embedding algebras) *)
11 (* Problem : Let g = cr(t). Show that t(tsg) = tt(ts)(tg) *)
13 (* Version : [Jec93] axioms : Incomplete > Reduced & Augmented > Incomplete. *)
17 (* Refs : [Jec93] Jech (1993), LD-Algebras *)
19 (* Source : [Jec93] *)
21 (* Names : Problem 8 [Jec93] *)
23 (* Status : Unsatisfiable *)
25 (* Rating : 0.00 v2.2.1, 0.11 v2.2.0, 0.14 v2.1.0, 0.13 v2.0.0 *)
27 (* Syntax : Number of clauses : 7 ( 0 non-Horn; 7 unit; 6 RR) *)
29 (* Number of atoms : 7 ( 7 equality) *)
31 (* Maximal clause size : 1 ( 1 average) *)
33 (* Number of predicates : 1 ( 0 propositional; 2-2 arity) *)
35 (* Number of functors : 9 ( 8 constant; 0-2 arity) *)
37 (* Number of variables : 3 ( 0 singleton) *)
39 (* Maximal term depth : 3 ( 2 average) *)
43 (* -------------------------------------------------------------------------- *)
45 (* ----Include Embedding algebra axioms *)
47 (* include('Axioms/LDA001-0.ax'). *)
49 (* -------------------------------------------------------------------------- *)
51 (* ----t(tsk) = tt(ts)(tk), where k=crit(t) *)
52 ntheorem prove_equation:
53 ∀Univ:Type.∀X:Univ.∀Y:Univ.∀Z:Univ.
54 ∀f:∀_:Univ.∀_:Univ.Univ.
63 ∀H0:eq Univ tsk (f ts k).
64 ∀H1:eq Univ tk (f t k).
65 ∀H2:eq Univ tt_ts (f tt ts).
66 ∀H3:eq Univ ts (f t s).
67 ∀H4:eq Univ tt (f t t).
68 ∀H5:∀X:Univ.∀Y:Univ.∀Z:Univ.eq Univ (f X (f Y Z)) (f (f X Y) (f X Z)).eq Univ (f t tsk) (f tt_ts tk)
89 nauto by H0,H1,H2,H3,H4,H5;
92 (* -------------------------------------------------------------------------- *)