1 include "logic/equality.ma".
3 (* Inclusion of: LAT078-1.p *)
5 (* -------------------------------------------------------------------------- *)
7 (* File : LAT078-1 : TPTP v3.7.0. Released v2.6.0. *)
9 (* Domain : Lattice Theory (Ortholattices) *)
11 (* Problem : Given single axiom MOL-27B2, prove associativity *)
13 (* Version : [MRV03] (equality) axioms. *)
15 (* English : Given a single axiom candidate MOL-27B2 for modular ortholattices *)
17 (* (MOL) in terms of the Sheffer Stroke, prove a Sheffer stroke form *)
19 (* of associativity. *)
21 (* Refs : [MRV03] McCune et al. (2003), Sheffer Stroke Bases for Ortholatt *)
23 (* Source : [MRV03] *)
25 (* Names : MOL-27B2-associativity [MRV03] *)
27 (* Status : Unsatisfiable *)
29 (* Rating : 1.00 v2.6.0 *)
31 (* Syntax : Number of clauses : 2 ( 0 non-Horn; 2 unit; 1 RR) *)
33 (* Number of atoms : 2 ( 2 equality) *)
35 (* Maximal clause size : 1 ( 1 average) *)
37 (* Number of predicates : 1 ( 0 propositional; 2-2 arity) *)
39 (* Number of functors : 4 ( 3 constant; 0-2 arity) *)
41 (* Number of variables : 4 ( 1 singleton) *)
43 (* Maximal term depth : 9 ( 4 average) *)
47 (* -------------------------------------------------------------------------- *)
49 (* ----Single axiom MOL-27B2 *)
51 (* ----Denial of Sheffer stroke associativity *)
52 ntheorem associativity:
53 (∀Univ:Type.∀A:Univ.∀B:Univ.∀C:Univ.∀D:Univ.
57 ∀f:∀_:Univ.∀_:Univ.Univ.
58 ∀H0:∀A:Univ.∀B:Univ.∀C:Univ.∀D:Univ.eq Univ (f (f (f (f B A) (f A C)) D) (f A (f (f (f B (f B (f (f C C) A))) A) C))) A.eq Univ (f a (f (f b c) (f b c))) (f c (f (f b a) (f b a))))
71 ntry (nassumption) ##;
74 (* -------------------------------------------------------------------------- *)