# HG changeset patch # User mengj # Date 1146197119 -7200 # Node ID 29c6cba140daa89af205c9e46ff914449e76b051 # Parent cd6c71c57f5371166f5e1402718f1a61f0626cfe added some helper files for HOL goals/lemmas. Clauses have TPTP format. diff -r cd6c71c57f53 -r 29c6cba140da src/HOL/Tools/atp-inputs/const_comb_inclS.tptp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/atp-inputs/const_comb_inclS.tptp Fri Apr 28 06:05:19 2006 +0200 @@ -0,0 +1,65 @@ +%ID: $Id$ +%Author: Jia Meng, NICTA +%typed combinator reduction for const-types-only (include S) + +%K P Q --> P +input_clause(a1,axiom, +[++equal(hAPP(hAPP(c_COMBK(A,B),P),Q),P)]). + +%S P Q R --> P R (Q R) +input_clause(a2,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBS(A,B,C),P),Q),R),hAPP(hAPP(P,R),hAPP(Q,R)))]). + + +%I P --> P +input_clause(a3,axiom, +[++equal(hAPP(c_COMBI(T),P),P)]). + +%B P Q R --> P(Q R) +input_clause(a4,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBB(A,B,C),P),Q),R),hAPP(P,hAPP(Q,R)))]). + + +%C P Q R --> P R Q +input_clause(a5,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBC(A,B,C),P),Q),R),hAPP(hAPP(P,R),Q))]). + +%the combinators are all different +input_clause(a6,axiom, +[--equal(c_COMBI(T),c_COMBK(A,B))]). + +input_clause(a7,axiom, +[--equal(c_COMBI(T),c_COMBS(A,B,C))]). + +input_clause(a8,axiom, +[--equal(c_COMBI(T),c_COMBB(A,B,C))]). + +input_clause(a9,axiom, +[--equal(c_COMBI(T),c_COMBC(A,B,C))]). + +input_clause(a10,axiom, +[--equal(c_COMBK(A,B),c_COMBS(A3,B3,C3))]). + +input_clause(a11,axiom, +[--equal(c_COMBK(A,B),c_COMBB(A1,B1,C1))]). + +input_clause(a12,axiom, +[--equal(c_COMBK(A,B),c_COMBC(A2,B2,C2))]). + +input_clause(a13,axiom, +[--equal(c_COMBS(A3,B3,C3),c_COMBB(A1,B1,C1))]). + +input_clause(a14,axiom, +[--equal(c_COMBS(A3,B3,C3),c_COMBC(A2,B2,C2))]). + +input_clause(a15,axiom, +[--equal(c_COMBB(A1,B1,C1),c_COMBC(A2,B2,C2))]). + +input_clause(a16,axiom, +[--hBOOL(hAPP(hAPP(fequal(A),X),Y)),++equal(X,Y)]). + +input_clause(a17,axiom, +[++hBOOL(hAPP(hAPP(fequal(A),X),Y)),--equal(X,Y)]). + + + diff -r cd6c71c57f53 -r 29c6cba140da src/HOL/Tools/atp-inputs/const_comb_noS.tptp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/atp-inputs/const_comb_noS.tptp Fri Apr 28 06:05:19 2006 +0200 @@ -0,0 +1,48 @@ +%ID: $Id$ +%Author: Jia Meng, NICTA +%typed combinator reduction for const-types-only (no S) + +%K P Q --> P +input_clause(a1,axiom, +[++equal(hAPP(hAPP(c_COMBK(A,B),P),Q),P)]). + +%I P --> P +input_clause(a3,axiom, +[++equal(hAPP(c_COMBI(T),P),P)]). + +%B P Q R --> P(Q R) +input_clause(a4,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBB(A,B,C),P),Q),R),hAPP(P,hAPP(Q,R)))]). + + +%C P Q R --> P R Q +input_clause(a5,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBC(A,B,C),P),Q),R),hAPP(hAPP(P,R),Q))]). + +%the combinators are all different +input_clause(a6,axiom, +[--equal(c_COMBI(T),c_COMBK(A,B))]). + +input_clause(a8,axiom, +[--equal(c_COMBI(T),c_COMBB(A,B,C))]). + +input_clause(a9,axiom, +[--equal(c_COMBI(T),c_COMBC(A,B,C))]). + +input_clause(a11,axiom, +[--equal(c_COMBK(A,B),c_COMBB(A1,B1,C1))]). + +input_clause(a12,axiom, +[--equal(c_COMBK(A,B),c_COMBC(A2,B2,C2))]). + +input_clause(a15,axiom, +[--equal(c_COMBB(A1,B1,C1),c_COMBC(A2,B2,C2))]). + +input_clause(a16,axiom, +[--hBOOL(hAPP(hAPP(fequal(A),X),Y)),++equal(X,Y)]). + +input_clause(a17,axiom, +[++hBOOL(hAPP(hAPP(fequal(A),X),Y)),--equal(X,Y)]). + + + diff -r cd6c71c57f53 -r 29c6cba140da src/HOL/Tools/atp-inputs/const_helper1.tptp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/atp-inputs/const_helper1.tptp Fri Apr 28 06:05:19 2006 +0200 @@ -0,0 +1,10 @@ +%ID: $Id$ +%Author: Jia Meng, NICTA +%extensionality for const-types-only + +input_clause(a18,axiom, +[--equal(hAPP(F,hAPP(hAPP(hEXTENT(A,B),F),G)),hAPP(G,hAPP(hAPP(hEXTENT(A,B),F),G))), +++equal(F,G)]). + + + diff -r cd6c71c57f53 -r 29c6cba140da src/HOL/Tools/atp-inputs/full_comb_inclS.tptp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/atp-inputs/full_comb_inclS.tptp Fri Apr 28 06:05:19 2006 +0200 @@ -0,0 +1,71 @@ +%ID: $Id$ +%Author: Jia Meng, NICTA +%typed combinator reduction for full-types (include S) + +%K P Q --> P +input_clause(a1,axiom, +[++equal(typeinfo(hAPP(typeinfo(hAPP(typeinfo(c_COMBK,tc_fun(A,tc_fun(B,A))),typeinfo(P,A)),tc_fun(B,A)),typeinfo(Q,B)),A),typeinfo(P,A))]). + +%S P Q R --> P R (Q R) +input_clause(a2,axiom, +[++equal(typeinfo(hAPP(typeinfo(hAPP(typeinfo(hAPP(typeinfo(c_COMBS,tc_fun(tc_fun(A,tc_fun(B,C)),tc_fun(tc_fun(A,B),tc_fun(A,C)))),typeinfo(P,tc_fun(A,tc_fun(B,C)))),tc_fun(tc_fun(A,B),tc_fun(A,C))),typeinfo(Q,tc_fun(A,B))),tc_fun(A,C)),typeinfo(R,A)),C),typeinfo(hAPP(typeinfo(hAPP(typeinfo(P,tc_fun(A,tc_fun(B,C))),typeinfo(R,A)),tc_fun(B,C)),typeinfo(hAPP(typeinfo(Q,tc_fun(A,B)),typeinfo(R,A)),B)),C))]). + +%I P --> P +input_clause(a3,axiom, +[++equal(typeinfo(hAPP(typeinfo(c_COMBI,tc_fun(T,T)),typeinfo(P,T)),T),typeinfo(P,T))]). + +%B P Q R --> P(Q R) +input_clause(a4,axiom, +[++equal(typeinfo(hAPP(typeinfo(hAPP(typeinfo(hAPP(typeinfo(c_COMBB,tc_fun(tc_fun(A,B),tc_fun(tc_fun(C,A),tc_fun(C,B)))),typeinfo(P,tc_fun(A,B))),tc_fun(tc_fun(C,A),tc_fun(C,B))),typeinfo(Q,tc_fun(C,A))),tc_fun(C,B)),typeinfo(R,C)),B),typeinfo(hAPP(typeinfo(P,tc_fun(A,B)),typeinfo(hAPP(typeinfo(Q,tc_fun(C,A)),typeinfo(R,C)),A)),B))]). + + +%C P Q R --> P R Q +input_clause(a5,axiom, +[++equal(typeinfo(hAPP(typeinfo(hAPP(typeinfo(hAPP(typeinfo(c_COMBC,tc_fun(tc_fun(A,tc_fun(B,C)),tc_fun(B,tc_fun(A,C)))),typeinfo(P,tc_fun(A,tc_fun(B,C)))),tc_fun(B,tc_fun(A,C))),typeinfo(Q,B)),tc_fun(A,C)),typeinfo(R,A)),C),typeinfo(hAPP(typeinfo(hAPP(typeinfo(P,tc_fun(A,tc_fun(B,C))),typeinfo(R,A)),tc_fun(B,C)),typeinfo(Q,B)),C))]). + +%the combinators are all different +input_clause(a6,axiom, +[--equal(typeinfo(c_COMBI,tc_fun(T,T)),typeinfo(c_COMBK,tc_fun(A,tc_fun(B,A))))]). + +input_clause(a7,axiom, +[--equal(typeinfo(c_COMBI,tc_fun(T,T)),typeinfo(c_COMBS,tc_fun(tc_fun(A,tc_fun(B,C)),tc_fun(tc_fun(A,B),tc_fun(A,C)))))]). + + +input_clause(a8,axiom, +[--equal(typeinfo(c_COMBI,tc_fun(T,T)),typeinfo(c_COMBB,tc_fun(tc_fun(A,B),tc_fun(tc_fun(C,A),tc_fun(C,B)))))]). + + +input_clause(a9,axiom, +[--equal(typeinfo(c_COMBI,tc_fun(T,T)),typeinfo(c_COMBC,tc_fun(tc_fun(A,tc_fun(B,C)),tc_fun(B,tc_fun(A,C)))))]). + + +input_clause(a10,axiom, +[--equal(typeinfo(c_COMBK,tc_fun(A,tc_fun(B,A))),typeinfo(c_COMBS,tc_fun(tc_fun(A3,tc_fun(B3,C3)),tc_fun(tc_fun(A3,B3),tc_fun(A3,C3)))))]). + +input_clause(a11,axiom, +[--equal(typeinfo(c_COMBK,tc_fun(A,tc_fun(B,A))),typeinfo(c_COMBB,tc_fun(tc_fun(A1,B1),tc_fun(tc_fun(C1,A1),tc_fun(C1,B1)))))]). + + +input_clause(a12,axiom, +[--equal(typeinfo(c_COMBK,tc_fun(A,tc_fun(B,A))),typeinfo(c_COMBC,tc_fun(tc_fun(A2,tc_fun(B2,C2)),tc_fun(B2,tc_fun(A2,C2)))))]). + + +input_clause(a13,axiom, +[--equal(typeinfo(c_COMBS,tc_fun(tc_fun(A3,tc_fun(B3,C3)),tc_fun(tc_fun(A3,B3),tc_fun(A3,C3)))),typeinfo(c_COMBB,tc_fun(tc_fun(A1,B1),tc_fun(tc_fun(C1,A1),tc_fun(C1,B1)))))]). + + +input_clause(a14,axiom, +[--equal(typeinfo(c_COMBS,tc_fun(tc_fun(A3,tc_fun(B3,C3)),tc_fun(tc_fun(A3,B3),tc_fun(A3,C3)))),typeinfo(c_COMBC,tc_fun(tc_fun(A2,tc_fun(B2,C2)),tc_fun(B2,tc_fun(A2,C2)))))]). + +input_clause(a15,axiom, +[--equal(typeinfo(c_COMBB,tc_fun(tc_fun(A1,B1),tc_fun(tc_fun(C1,A1),tc_fun(C1,B1)))),typeinfo(c_COMBC,tc_fun(tc_fun(A2,tc_fun(B2,C2)),tc_fun(B2,tc_fun(A2,C2)))))]). + +input_clause(a16,axiom, +[--hBOOL(typeinfo(hAPP(typeinfo(hAPP(typeinfo(fequal,tc_fun(A,tc_fun(A,tc_bool))),typeinfo(X,A)),tc_fun(A,tc_bool)),typeinfo(Y,A)),tc_bool)), +++equal(typeinfo(X,A),typeinfo(Y,A))]). + +input_clause(a17,axiom, +[++hBOOL(typeinfo(hAPP(typeinfo(hAPP(typeinfo(fequal,tc_fun(A,tc_fun(A,tc_bool))),typeinfo(X,A)),tc_fun(A,tc_bool)),typeinfo(Y,A)),tc_bool)), +--equal(typeinfo(X,A),typeinfo(Y,A))]). + + diff -r cd6c71c57f53 -r 29c6cba140da src/HOL/Tools/atp-inputs/full_comb_noS.tptp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/atp-inputs/full_comb_noS.tptp Fri Apr 28 06:05:19 2006 +0200 @@ -0,0 +1,52 @@ +%ID: $Id$ +%Author: Jia Meng, NICTA +%typed combinator reduction for full-types (no S) + +%K P Q --> P +input_clause(a1,axiom, +[++equal(typeinfo(hAPP(typeinfo(hAPP(typeinfo(c_COMBK,tc_fun(A,tc_fun(B,A))),typeinfo(P,A)),tc_fun(B,A)),typeinfo(Q,B)),A),typeinfo(P,A))]). + +%I P --> P +input_clause(a3,axiom, +[++equal(typeinfo(hAPP(typeinfo(c_COMBI,tc_fun(T,T)),typeinfo(P,T)),T),typeinfo(P,T))]). + +%B P Q R --> P(Q R) +input_clause(a4,axiom, +[++equal(typeinfo(hAPP(typeinfo(hAPP(typeinfo(hAPP(typeinfo(c_COMBB,tc_fun(tc_fun(A,B),tc_fun(tc_fun(C,A),tc_fun(C,B)))),typeinfo(P,tc_fun(A,B))),tc_fun(tc_fun(C,A),tc_fun(C,B))),typeinfo(Q,tc_fun(C,A))),tc_fun(C,B)),typeinfo(R,C)),B),typeinfo(hAPP(typeinfo(P,tc_fun(A,B)),typeinfo(hAPP(typeinfo(Q,tc_fun(C,A)),typeinfo(R,C)),A)),B))]). + + +%C P Q R --> P R Q +input_clause(a5,axiom, +[++equal(typeinfo(hAPP(typeinfo(hAPP(typeinfo(hAPP(typeinfo(c_COMBC,tc_fun(tc_fun(A,tc_fun(B,C)),tc_fun(B,tc_fun(A,C)))),typeinfo(P,tc_fun(A,tc_fun(B,C)))),tc_fun(B,tc_fun(A,C))),typeinfo(Q,B)),tc_fun(A,C)),typeinfo(R,A)),C),typeinfo(hAPP(typeinfo(hAPP(typeinfo(P,tc_fun(A,tc_fun(B,C))),typeinfo(R,A)),tc_fun(B,C)),typeinfo(Q,B)),C))]). + +%the combinators are all different +input_clause(a6,axiom, +[--equal(typeinfo(c_COMBI,tc_fun(T,T)),typeinfo(c_COMBK,tc_fun(A,tc_fun(B,A))))]). + +input_clause(a8,axiom, +[--equal(typeinfo(c_COMBI,tc_fun(T,T)),typeinfo(c_COMBB,tc_fun(tc_fun(A,B),tc_fun(tc_fun(C,A),tc_fun(C,B)))))]). + + +input_clause(a9,axiom, +[--equal(typeinfo(c_COMBI,tc_fun(T,T)),typeinfo(c_COMBC,tc_fun(tc_fun(A,tc_fun(B,C)),tc_fun(B,tc_fun(A,C)))))]). + +input_clause(a11,axiom, +[--equal(typeinfo(c_COMBK,tc_fun(A,tc_fun(B,A))),typeinfo(c_COMBB,tc_fun(tc_fun(A1,B1),tc_fun(tc_fun(C1,A1),tc_fun(C1,B1)))))]). + + +input_clause(a12,axiom, +[--equal(typeinfo(c_COMBK,tc_fun(A,tc_fun(B,A))),typeinfo(c_COMBC,tc_fun(tc_fun(A2,tc_fun(B2,C2)),tc_fun(B2,tc_fun(A2,C2)))))]). + + +input_clause(a15,axiom, +[--equal(typeinfo(c_COMBB,tc_fun(tc_fun(A1,B1),tc_fun(tc_fun(C1,A1),tc_fun(C1,B1)))),typeinfo(c_COMBC,tc_fun(tc_fun(A2,tc_fun(B2,C2)),tc_fun(B2,tc_fun(A2,C2)))))]). + +input_clause(a16,axiom, +[--hBOOL(typeinfo(hAPP(typeinfo(hAPP(typeinfo(fequal,tc_fun(A,tc_fun(A,tc_bool))),typeinfo(X,A)),tc_fun(A,tc_bool)),typeinfo(Y,A)),tc_bool)), +++equal(typeinfo(X,A),typeinfo(Y,A))]). + +input_clause(a17,axiom, +[++hBOOL(typeinfo(hAPP(typeinfo(hAPP(typeinfo(fequal,tc_fun(A,tc_fun(A,tc_bool))),typeinfo(X,A)),tc_fun(A,tc_bool)),typeinfo(Y,A)),tc_bool)), +--equal(typeinfo(X,A),typeinfo(Y,A))]). + + diff -r cd6c71c57f53 -r 29c6cba140da src/HOL/Tools/atp-inputs/full_helper1.tptp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/atp-inputs/full_helper1.tptp Fri Apr 28 06:05:19 2006 +0200 @@ -0,0 +1,7 @@ +%ID: $Id$ +%Author: Jia Meng, NICTA +%extensionality for full-types + +input_clause(a18,axiom, +[--equal(typeinfo(hAPP(typeinfo(F,tc_fun(A,B)),typeinfo(hAPP(typeinfo(hAPP(typeinfo(hEXTENT,tc_fun(tc_fun(A,B),tc_fun(tc_fun(A,B),A))),typeinfo(F,tc_fun(A,B))),tc_fun(tc_fun(A,B),A)),typeinfo(G,tc_fun(A,B))),A)),B),typeinfo(hAPP(typeinfo(G,tc_fun(A,B)),typeinfo(hAPP(typeinfo(hAPP(typeinfo(hEXTENT,tc_fun(tc_fun(A,B),tc_fun(tc_fun(A,B),A))),typeinfo(F,tc_fun(A,B))),tc_fun(tc_fun(A,B),A)),typeinfo(G,tc_fun(A,B))),A)),B)), +++equal(typeinfo(F,tc_fun(A,B)),typeinfo(G,tc_fun(A,B)))]). diff -r cd6c71c57f53 -r 29c6cba140da src/HOL/Tools/atp-inputs/par_comb_inclS.tptp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/atp-inputs/par_comb_inclS.tptp Fri Apr 28 06:05:19 2006 +0200 @@ -0,0 +1,67 @@ +%ID: $Id$ +%Author: Jia Meng, NICTA +%typed combinator reduction for partial-types (include S) + +%K P Q --> P +input_clause(a1,axiom, +[++equal(hAPP(hAPP(c_COMBK,P,tc_fun(A,tc_fun(B,A))),Q,tc_fun(B,A)),P)]). + + +%S P Q R --> P R (Q R) +input_clause(a2,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBS,P,tc_fun(tc_fun(A,tc_fun(B,C)),tc_fun(tc_fun(A,B),tc_fun(A,C)))),Q,tc_fun(tc_fun(A,B),tc_fun(A,C))),R,tc_fun(A,C)),hAPP(hAPP(P,R,tc_fun(A,tc_fun(B,C))),hAPP(Q,R,tc_fun(A,B)),tc_fun(B,C)))]). + +%I P --> P +input_clause(a3,axiom, +[++equal(hAPP(c_COMBI,P,tc_fun(T,T)),P)]). + + +%B P Q R --> P(Q R) +input_clause(a4,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBB,P,tc_fun(tc_fun(A,B),tc_fun(tc_fun(C,A),tc_fun(C,B)))),Q,tc_fun(tc_fun(C,A),tc_fun(C,B))),R,tc_fun(C,B)),hAPP(P,hAPP(Q,R,tc_fun(C,A)),tc_fun(A,B)))]). + + +%C P Q R --> P R Q +input_clause(a5,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBC,P,tc_fun(tc_fun(A,tc_fun(B,C)),tc_fun(B,tc_fun(A,C)))),Q,tc_fun(B,tc_fun(A,C))),R,tc_fun(A,C)),hAPP(hAPP(P,R,tc_fun(A,tc_fun(B,C))),Q,tc_fun(B,C)))]). + +%the combinators are all different +input_clause(a6,axiom, +[--equal(c_COMBI,c_COMBK)]). + +input_clause(a7,axiom, +[--equal(c_COMBI,c_COMBS)]). + +input_clause(a8,axiom, +[--equal(c_COMBI,c_COMBB)]). + +input_clause(a9,axiom, +[--equal(c_COMBI,c_COMBC)]). + +input_clause(a10,axiom, +[--equal(c_COMBK,c_COMBS)]). + +input_clause(a11,axiom, +[--equal(c_COMBK,c_COMBB)]). + +input_clause(a12,axiom, +[--equal(c_COMBK,c_COMBC)]). + +input_clause(a13,axiom, +[--equal(c_COMBS,c_COMBB)]). + +input_clause(a14,axiom, +[--equal(c_COMBS,c_COMBC)]). + +input_clause(a15,axiom, +[--equal(c_COMBB,c_COMBC)]). + +input_clause(a16,axiom, +[--hBOOL(hAPP(hAPP(fequal,X,tc_fun(A,tc_fun(A,tc_bool))),Y,tc_fun(A,tc_bool))), +++equal(X,Y)]). + +input_clause(a17,axiom, +[++hBOOL(hAPP(hAPP(fequal,X,tc_fun(A,tc_fun(A,tc_bool))),Y,tc_fun(A,tc_bool))), +--equal(X,Y)]). + + diff -r cd6c71c57f53 -r 29c6cba140da src/HOL/Tools/atp-inputs/par_comb_noS.tptp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/atp-inputs/par_comb_noS.tptp Fri Apr 28 06:05:19 2006 +0200 @@ -0,0 +1,50 @@ +%ID: $Id$ +%Author: Jia Meng, NICTA +%typed combinator reduction for partial-types (no S) + +%K P Q --> P +input_clause(a1,axiom, +[++equal(hAPP(hAPP(c_COMBK,P,tc_fun(A,tc_fun(B,A))),Q,tc_fun(B,A)),P)]). + +%I P --> P +input_clause(a3,axiom, +[++equal(hAPP(c_COMBI,P,tc_fun(T,T)),P)]). + + +%B P Q R --> P(Q R) +input_clause(a4,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBB,P,tc_fun(tc_fun(A,B),tc_fun(tc_fun(C,A),tc_fun(C,B)))),Q,tc_fun(tc_fun(C,A),tc_fun(C,B))),R,tc_fun(C,B)),hAPP(P,hAPP(Q,R,tc_fun(C,A)),tc_fun(A,B)))]). + + +%C P Q R --> P R Q +input_clause(a5,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBC,P,tc_fun(tc_fun(A,tc_fun(B,C)),tc_fun(B,tc_fun(A,C)))),Q,tc_fun(B,tc_fun(A,C))),R,tc_fun(A,C)),hAPP(hAPP(P,R,tc_fun(A,tc_fun(B,C))),Q,tc_fun(B,C)))]). + +%the combinators are all different +input_clause(a6,axiom, +[--equal(c_COMBI,c_COMBK)]). + +input_clause(a8,axiom, +[--equal(c_COMBI,c_COMBB)]). + +input_clause(a9,axiom, +[--equal(c_COMBI,c_COMBC)]). + +input_clause(a11,axiom, +[--equal(c_COMBK,c_COMBB)]). + +input_clause(a12,axiom, +[--equal(c_COMBK,c_COMBC)]). + +input_clause(a15,axiom, +[--equal(c_COMBB,c_COMBC)]). + +input_clause(a16,axiom, +[--hBOOL(hAPP(hAPP(fequal,X,tc_fun(A,tc_fun(A,tc_bool))),Y,tc_fun(A,tc_bool))), +++equal(X,Y)]). + +input_clause(a17,axiom, +[++hBOOL(hAPP(hAPP(fequal,X,tc_fun(A,tc_fun(A,tc_bool))),Y,tc_fun(A,tc_bool))), +--equal(X,Y)]). + + diff -r cd6c71c57f53 -r 29c6cba140da src/HOL/Tools/atp-inputs/par_helper1.tptp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/atp-inputs/par_helper1.tptp Fri Apr 28 06:05:19 2006 +0200 @@ -0,0 +1,9 @@ +%ID: $Id$ +%Author: Jia Meng, NICTA +%extensionality for partial types + +input_clause(a18,axiom, +[--equal(hAPP(F,hAPP(hAPP(hEXTENT,F,tc_fun(tc_fun(A,B),tc_fun(tc_fun(A,B),A))),G,tc_fun(tc_fun(A,B),A)),tc_fun(A,B)),hAPP(G,hAPP(hAPP(hEXTENT,F,tc_fun(tc_fun(A,B),tc_fun(tc_fun(A,B),A))),G,tc_fun(tc_fun(A,B),A)),tc_fun(A,B))), +++equal(F,G)]). + + diff -r cd6c71c57f53 -r 29c6cba140da src/HOL/Tools/atp-inputs/u_comb_inclS.tptp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/atp-inputs/u_comb_inclS.tptp Fri Apr 28 06:05:19 2006 +0200 @@ -0,0 +1,61 @@ +%ID: $Id$ +%Author: Jia Meng, NICTA +%untyped combinator reduction (include S) + +input_clause(a1,axiom, +[++equal(hAPP(hAPP(c_COMBK,P),Q),P)]). + +%S P Q R +input_clause(a2,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBS,P),Q),R),hAPP(hAPP(P,R),hAPP(Q,R)))]). + +input_clause(a3,axiom, +[++equal(hAPP(c_COMBI,P),P)]). + +%B P Q R --> P(Q R) +input_clause(a4,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBB,P),Q),R),hAPP(P,hAPP(Q,R)))]). + +%C P Q R --> P R Q +input_clause(a5,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBC,P),Q),R),hAPP(hAPP(P,R),Q))]). + +%the combinators are all different +input_clause(a6,axiom, +[--equal(c_COMBI,c_COMBK)]). + +input_clause(a7,axiom, +[--equal(c_COMBI,c_COMBS)]). + +input_clause(a8,axiom, +[--equal(c_COMBI,c_COMBB)]). + +input_clause(a9,axiom, +[--equal(c_COMBI,c_COMBC)]). + +input_clause(a10,axiom, +[--equal(c_COMBK,c_COMBS)]). + +input_clause(a11,axiom, +[--equal(c_COMBK,c_COMBB)]). + +input_clause(a12,axiom, +[--equal(c_COMBK,c_COMBC)]). + +input_clause(a13,axiom, +[--equal(c_COMBS,c_COMBB)]). + +input_clause(a14,axiom, +[--equal(c_COMBS,c_COMBC)]). + +input_clause(a15,axiom, +[--equal(c_COMBB,c_COMBC)]). + +input_clause(a16,axiom, +[--hBOOL(hAPP(hAPP(fequal,X),Y)),++equal(X,Y)]). + +input_clause(a17,axiom, +[++hBOOL(hAPP(hAPP(fequal,X),Y)),--equal(X,Y)]). + + + diff -r cd6c71c57f53 -r 29c6cba140da src/HOL/Tools/atp-inputs/u_comb_noS.tptp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/atp-inputs/u_comb_noS.tptp Fri Apr 28 06:05:19 2006 +0200 @@ -0,0 +1,45 @@ +%ID: $Id$ +%Author: Jia Meng, NICTA +%untyped combinator reduction (no S) + +input_clause(a1,axiom, +[++equal(hAPP(hAPP(c_COMBK,P),Q),P)]). + +input_clause(a3,axiom, +[++equal(hAPP(c_COMBI,P),P)]). + +%B P Q R --> P(Q R) +input_clause(a4,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBB,P),Q),R),hAPP(P,hAPP(Q,R)))]). + +%C P Q R --> P R Q +input_clause(a5,axiom, +[++equal(hAPP(hAPP(hAPP(c_COMBC,P),Q),R),hAPP(hAPP(P,R),Q))]). + +%the combinators are all different +input_clause(a6,axiom, +[--equal(c_COMBI,c_COMBK)]). + +input_clause(a8,axiom, +[--equal(c_COMBI,c_COMBB)]). + +input_clause(a9,axiom, +[--equal(c_COMBI,c_COMBC)]). + +input_clause(a11,axiom, +[--equal(c_COMBK,c_COMBB)]). + +input_clause(a12,axiom, +[--equal(c_COMBK,c_COMBC)]). + +input_clause(a15,axiom, +[--equal(c_COMBB,c_COMBC)]). + +input_clause(a16,axiom, +[--hBOOL(hAPP(hAPP(fequal,X),Y)),++equal(X,Y)]). + +input_clause(a17,axiom, +[++hBOOL(hAPP(hAPP(fequal,X),Y)),--equal(X,Y)]). + + + diff -r cd6c71c57f53 -r 29c6cba140da src/HOL/Tools/atp-inputs/u_helper1.tptp --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/atp-inputs/u_helper1.tptp Fri Apr 28 06:05:19 2006 +0200 @@ -0,0 +1,8 @@ +%ID: $Id$ +%Author: Jia Meng, NICTA +%extensionality, untyped + +input_clause(a18,axiom, +[--equal(hAPP(F,hAPP(hAPP(hEXTENT,F),G)),hAPP(G,hAPP(hAPP(hEXTENT,F),G))), +++equal(F,G)]). +