Prover/index.html
changeset 96 907b1fff5637
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Prover/index.html	Thu Mar 15 10:07:28 2012 +0000
@@ -0,0 +1,384 @@
+<HTML>
+<HEAD>
+   <META HTTP-EQUIV="Content-Type" CONTENT="text/html; charset=iso-8859-1">
+   <TITLE>G4ip</TITLE>
+</HEAD>
+<BODY TEXT="#000000" BGCOLOR="#C7C3C7" LINK="#0000EF" VLINK="#51188E" ALINK="#FF0000">
+
+<H2>An Implementation of G4ip in Pizza</H2>
+
+<FONT COLOR="#800000"><B>Warning:</B></FONT> 
+This page is now rather old! While you might still be interested
+in the algorithms, Robert Macdonald reported that Pizza and the current Java 
+implementation (version 1.3.0) do not work together. This means you need to 
+install an older Java version if you want to recompile the files given below. 
+I am happy to answer all question concerning the prover, but be aware that 
+currently for any kind of Java stuff I am using MLJ, which as of writing 
+this note  has not yet been made available for the general audience (maybe
+in the future also its OCaml equivalent). So I am not very fluent in Pizza
+anymore. <B>Update</B> Pizza development is continued and starting from version 
+<A HREF="http://pizzacompiler.sourceforge.net/">0.40</A> it should work
+with recent Java implementations.<P> 
+
+
+Jump to the <A HREF="#Implementation">implementation.</a>
+
+
+<H4>Introduction</H4>
+
+
+A convenient representation of intuitionistic logic is Getzen's 
+sequent calculus LJ (also G1i). A sequent of LJ can be proved 
+by applying inference rules until one reaches axioms, or can make no further 
+progress in which case one must backtrack or even abandon the search. 
+Unfortunately an interpreter for LJ using this depth-first strategy cannot 
+guarantee termination of the proof search.  Several modifications can be 
+made to LJ's inference rules without loss of soundness and completeness.  
+As result an efficient depth-first proof search can be designed for the 
+propositional fragment of intuitionistic logic. The name G4ip has been 
+assigned to the corresponding calculus in 
+<A HREF="#TroelstraSchwichtenberg96">[Troelstra and Schwichtenberg, 1996]</a>. 
+This calculus is also known as LJT which has been studied thoroughly 
+in <A HREF="#Dyckhoff92">[Dyckhoff, 1992]</a>. The inference rules of 
+G4ip are given <A HREF="G4ip.html">here</A>.<P>
+
+It is not very complicated to implement an interpreter for G4ip using a logic 
+programming language (backtracking is directly supported within the language).
+Our first implementation is written in the logic programming language
+<A HREF="http://www.cis.upenn.edu/~dale/lProlog/terzo/index.html">Lambda Prolog</A> 
+and can be found <A HREF="G4ip.mod">here</A>. Another implementation by
+Hodas and Miller written in <A HREF="http://www.cs.hmc.edu/~hodas/research/lolli/">Lolli</a>
+can be found <A HREF="ftp://ftp.cse.psu.edu/pub/dale/ic94-code/index.html">here</A>
+(see <A HREF="#HodasMiller94">[Hodas and Miller, 1994]</a>). These are simple and 
+straightforward implementations of G4ip's rules. On the other hand it seems 
+that imperative languages need a rather high overhead of code when implementing
+a logic calculus. For example choice points are usually implemented with stacks.
+We shall demonstrate the implementation technique of success 
+continuations which provides an equally simple method for implementing logic calculi 
+in imperative languages. This technique is not new: it has been introduced in 
+<A HREF="#Carlsson84">[Carlsson, 1984]</a>. This paper presents a rather technical
+implementation of Prolog in LISP. Later an excellent paper, 
+<A HREF="#ElliotPfenning91">[Elliot and Pfenning, 1991]</a>, appeared which 
+describes a full-fledged implementation of Lambda Prolog in SML.
+We demonstrate the technique of success continuations for G4ip in 
+<A HREF="http://www.cis.unisa.edu.au/~pizza/">Pizza</A>.<P>
+
+Pizza is an object-oriented programming language and an attractive extension of  
+<A HREF="http://www.javasoft.com/">Java</A>. Although Pizza is a superset of
+Java, Pizza programs can be translated into Java or compiled into ordinary
+Java Byte Code (see <A HREF="#OderskyWadler97">[Odersky and Wadler, 1997]</a>
+for a technical introduction to Pizza). We make use of the following two new 
+features of Pizza:
+<UL>
+<LI> higher-order functions, i.e. functions may be passed as parameters or returned
+from methods,
+<LI> class cases and pattern matching: this allows much simpler and more readable code.
+</UL>
+
+These features are not directly present in Java, but Pizza makes them accessible by 
+translating them into Java. Pizza provides the programmer with the same 
+extensive libraries for graphic and network applications as Java. The higher-order 
+functions are essential for the technique of success continuations. The success 
+continuations are functions passed as parameters or returned as values.<BR>
+
+
+<H4>The Representation of Formulae and Sequents</H4>
+
+Amongst the new language features of Pizza are class cases and pattern
+matching, which provide a very pleasant syntax for algebraic data types. The
+formulae of G4ip are specified by the following grammar:<P>
+
+<CODE>F ::= false | A | F & F | F v F | F -> F</CODE><P>
+
+
+The class cases allow a straightforward implementation of this specification;
+it is analogous to the SML implementation of 
+<A HREF="http://www.cis.upenn.edu/~dale/lProlog/terzo/index.html">Lambda Prolog's</A> 
+formulae in <A HREF="#ElliotPfenning91">[Elliot and Pfenning, 1991]</A>. The class 
+of formulae for G4ip is given below:<P>
+
+<TT>
+<DL>
+<DD>public class Form {            </DD>
+<DD>&nbsp;&nbsp;    case False();              </DD>    
+<DD>&nbsp;&nbsp;    case Atm(String c);        </DD>  
+<DD>&nbsp;&nbsp;    case And(Form c1,Form c2); </DD>
+<DD>&nbsp;&nbsp;    case Or(Form c1,Form c2);  </DD>
+<DD>&nbsp;&nbsp;    case Imp(Form c1,Form c2); </DD>
+<DD>}                              </DD>
+</DL>
+</TT>
+
+Two examples that illustrate the use of the representation are as follows:<P>
+
+
+<CODE>&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;p -> p &nbsp;</CODE> 
+is represented as   <CODE> &nbsp; Imp(Atm("p"),Atm("p"))</CODE><BR>
+<CODE>a v (a -> false) &nbsp;</CODE> is represented as  <CODE> &nbsp; Or(Atm("a"),Imp(Atm("a"),False()))</CODE><P>
+
+The class cases of Pizza also support an implementation of formulae specified
+by a mutually recursive grammar.  This is required, for example, when
+implementing hereditary Harrop formulae.<P> 
+
+The sequents of G4ip, which have the form <CODE>Gamma=>G</CODE>, are represented
+by means of the class below. The left-hand side of each sequent is specified by a multiset 
+of formulae.  Therefore, we do not need to worry about the order in which the
+formulae occur.<P>
+
+<TT>
+<DL>
+<DD>public class Sequent  {                       </DD>
+<DD>&nbsp;&nbsp; Form  G;                                      </DD>         
+<DD>&nbsp;&nbsp; Context Gamma;                                </DD>
+<DD>&nbsp;&nbsp; public Sequent(Context _Gamma, Form _G) {...};</DD>
+<DD>}                                             </DD>
+</DL>
+</TT>
+
+We have a constructor for generating new sequents during proof search. 
+<CODE>Context</CODE> is a class which represents multisets; it is a simple 
+extension of the class <CODe>Vector</CODE> available in the Java libraries. 
+This class provides methods for adding elements to a multiset (<CODE>add</CODE>), 
+taking out elements from a multiset (<CODE>removeElement</CODE>) and testing
+the membership of an element in a multiset (<CODE>includes</CODE>).
+
+
+<H4>The Technique of Success Continuations</H4>
+
+We have to distinguish between the concepts of proof obligations (which must 
+be proved) and choice points (which can be tried out to construct a proof). 
+The first argument of the method <CODE>prove</CODE> is the sequent being 
+proved; the second argument is an anonymous function. The function <CODE>prove</CODE> is now 
+of the form <CODE>prove(sequent,sc)</CODE>.  Somewhat simplified the
+first argument is the leftmost premise and the second argument <CODE>sc</CODE>, 
+the success continuation, represents the other proof obligations. In case we
+succeed in proving the first premise we then can attempt to prove the other
+premises. The technique of success continuations will be explained using the following 
+proof (each sequent is marked with a number):<P>
+
+<UL><img src="proof.gif" width=337 height=112></UL>
+<BR><P>
+
+The inference rules fall into three groups:
+
+<UL>
+<LI>  inference rules with a single premise (e.g. ->_R, &_L), 
+<LI>  inference rules with two premises (e.g. v_L) and
+<LI>  inference rules without premises (e.g. Axiom). 
+</UL>
+
+The following picture shows the order in which the sequents are being proved.
+
+<UL><img src="execution.gif" width=358 height=191></UL>
+<BR><P>
+
+Suppose we have called <CODE>prove</CODE> with a sequent <B>s</B> and a 
+success continuation <B>is</B>. The inference rules of the first
+group manipulate  <B>s</B> obtaining <B>s'</B> and call <CODE>prove</CODE> 
+again with the new sequent <B>s'</B> and the current success continuation 
+(Steps 1-2, 3-4 and 5-6).  The inference rules 
+of the second group have two premises, <B>s1</B> and <B>s2</B>. 
+These rules call <CODE>prove</CODE> with <B>s1</B> and a new success 
+continuation <CODE>prove(s2,is)</CODE> (Step 2-3).
+The third group of inference rules only invoke the success continuation 
+if the rule was applicable (Steps 4-5 and 6-7).<P>
+
+
+We are going to give a detailed description of the code for the rules: &_L,
+->_R, v_Ri, v_L and Axiom.  The function <CODE>prove</CODE> receives as arguments 
+a sequent <CODE>Sequent(Gamma,G)</CODE> and a success continuation 
+<CODE>sc</CODE>.  It enumerates all formulae as being principal and 
+two switch statements select a corresponding case depending on the form 
+and the occurrence of the principal formula.<P> 
+
+The &_L rule is in the first group; it modifies the sequent being proved and calls
+<CODE>prove</CODE> again with the current success continuation sc. The code is as
+follows (<CODE>Gamma</CODE> stands for the set of formulae on the left-hand
+side of a sequent excluding the principal formula; <CODE>G</CODE> stands 
+for the goal formula of a sequent; <CODE>B</CODE> and <CODE>C</CODE> stand 
+for the two components of the principal formula).<P>
+
+<TT>
+<DL>
+<DD>case And(Form B, Form C):</DD>
+<DD>&nbsp;&nbsp; prove(new Sequent(Gamma.add(B,C),G),sc); break;</DD>
+</DL>
+</TT>
+
+The code for the ->_R rule is similar:<P>
+
+<TT>
+<DL>
+<DD>case Imp(Form B, Form C):</DD>
+<DD>&nbsp;&nbsp; prove(new Sequent(Gamma.add(A),B),sc); break;</DD>
+</DL>
+</TT>
+
+The v_Ri rule is an exception in the first group. It breaks up a goal
+formula of the form <CODE>B1 v B2</CODE> and proceeds with one of its component. 
+Since we do not know in advance which component leads to a successful proof we have
+to try both. Therefore this rule acts as a choice point, which is encoded by a
+recursive call of <CODE>prove</CODE> for each case.
+
+<TT>
+<DL>
+<DD>case Or(Form B1,Form B2):</DD>
+<DD>&nbsp;&nbsp; prove(new Sequent(Gamma,B1),sc);</DD>
+<DD>&nbsp;&nbsp; prove(new Sequent(Gamma,B2),sc); break;</DD>
+</DL>
+</TT>
+
+The v_L rule falls into the second group where the current success
+continuation, sc, is modified. It calls <CODE>prove</CODE> with the first premise,
+<CODE>B,Gamma=>G</CODE>, and wraps up the success continuation with the 
+new proof obligation, <CODE>C,Gamma=>G</CODE>. The construction 
+<CODE>fun()->void {...}</CODE> defines an anonymous function: the new 
+success continuation. In case the sequent <CODE>B,Gamma=>G</CODE> can be 
+proved, this function is invoked.
+
+<TT>
+<DL>
+<DD>case Or(Form B,Form C):</DD>
+<DD>&nbsp;&nbsp; prove(new Sequent(Gamma.add(B),G),</DD>
+<DD>&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;&nbsp; 
+      fun()->void {prove(new Sequent(Gamma.add(C),G),sc);}</DD>
+<DD>&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;&nbsp;&nbsp; ); break</DD>
+</DL>
+</TT>
+
+The Axiom rule falls into the third group. It first checks if the
+principal formula (which is an atom) matches with the goal formula and
+then invokes the success continuation sc in order to prove all remaining
+proof obligations.
+
+<TT>
+<DL>
+<DD>case Atm(String c):</DD>
+<DD>&nbsp;&nbsp; if (G instanceof Atm) { </DD>
+<DD>&nbsp;&nbsp; &nbsp;&nbsp; if (G.c.compareTo(c) == 0) { sc(); }</DD>
+<DD>&nbsp;&nbsp; } break;</DD>
+</DL>
+</TT>
+
+The proof search is started with an initial success continuation <B>is</B>. 
+This initial success continuation is invoked when a proof has been found. 
+In this case we want to give some response to the user, an
+example for the initial success continuation could be as follows:
+
+<TT>
+<DL>
+<DD> public void initial_sc()  { System.out.println("Provable!"); } </DD>
+</DL>
+</TT>
+
+
+Suppose we attempt to start the proof search with <CODE>prove(p,p => p,is)</CODE>. 
+We would find that the prover responds twice with <CODE>"Provable!"</CODE>, because 
+it finds two proofs. In our implementation this problem is avoided by encoding 
+the proof search as a thread. Whenever a proof is found, the initial success 
+continuation displays the proof and suspends the thread. The user can 
+decide to resume with the proof search or abandon the search.
+
+
+<H4>Conclusion</H4>
+
+The implementation cannot be considered as optimal in terms of speed. 
+A much more efficient algorithm for G4ip (but less clear) has been 
+implemented by Dyckhoff in Prolog. Similar ideas can be encoded in our 
+Pizza implementation; but our point was not the efficiency but the clarity 
+of the implementation using success continuations.
+The technique is applicable elsewhere whenever backtracking is required.  We
+compared the code of our implementation with an implementation in
+<A HREF="http://www.cis.upenn.edu/~dale/lProlog/terzo/index.html">Lambda Prolog</A>: 
+the ratio of code is approximately 2 to 1. 
+(see <A HREF="G4ip.mod">LambdaProlog code</A> and 
+<A HREF="minimal/MinProver.pizza">Pizza code</A>).
+This result is partly due to the fact that we had to implement a class for 
+multisets. In a future version of Java, we could have accessed a package 
+in the library. The technique of success continuation can also be applied 
+to a first-order calculus as shown in <A HREF="#ElliotPfenning91">[Elliot and Pfenning, 1991]</a>, 
+but the required mechanism of substitution needs to be implemented separately. 
+However, we think the technique of success continuations provides a remarkable 
+simple implementation for logic calculi.<P> 
+
+We had to make some compromises in order to support as many platforms
+as possible. This should change with the release of new browsers and a stable
+Java-specification (resp. Pizza-specification).<P> 
+
+A paper about the implementation appeared in the LNAI series No 1397,
+Automated Reasoning with Analytic Tableaux and Related Methods, 
+ed. Harry de Swart, International Conference Tableaux'98 in Oisterwijk, 
+The Netherlands. The title is: Implementation of Proof Search in 
+the Imperative Programming Language Pizza (pp. 313-319). The paper can be 
+found here: <A HREF="Tableaux98.dvi.gz">DVI</A>, <A HREF="Tableaux98.ps.gz">Postscript</A> 
+(© Springer-Verlag <A HREF="http://www.springer.de/comp/lncs/index.html">LNCS</A>).<P>
+
+
+<B>Acknowledgements:</B> I am very grateful for Dr Roy Dyckhoff's constant
+encouragement and many comments on my work. I thank Dr Gavin Bierman who
+helped me to test the prover applet. 
+
+<HR>
+<A NAME="Implementation"></A><H4>Implementation</H4>
+
+<A HREF="README">Readme</A><p>
+
+<A HREF="ProverApplet.html"><B>Prover Applet</B></A><BR>
+<A HREF="ProverAppletJar.html"><B>Jar Version</B></A> 
+(slightly faster, but requires Netscape 4 or MS Explorer 4).<P>
+
+
+<HR>
+<B>References</B> 
+<UL>
+<LI> <A NAME="Carlsson84"></A>
+     [Carlsson, 1984]<BR> 
+     M. Carlsson, On  Implementing Prolog in Functional Programming, 
+     New Generation Computing, pp 347-359.
+<LI> <A NAME="Dyckhoff92"></A>
+     [Dyckhoff, 1992]<BR> 
+     <A HREF="http://www-theory.dcs.st-and.ac.uk/~rd/">R. Dyckhoff</A>, 
+     Contraction-Free Sequent Calculi for Intuitionistic Logic, 
+     Journal of Symbolic Logic 57(3), pp 795-807.
+<LI> <A NAME="ElliotPfenning91"></A>
+     [Elliot and Pfenning, 1991]<BR> 
+     C. Elliot,
+     <A HREF="http://foxnet.cs.cmu.edu/people/fp/homepage.html">F. Pfenning</A>, 
+     A Semi-Functional Implementation of a Higher-Order Programming Language, 
+     In Peter Lee, editor, Topics in Advanced Language Implementation, MIT Press,
+     pp 289-352. 
+     <A HREF="http://www.cs.cmu.edu/~fp/papers/elpsml-paper.tar.gz">Available electronically</a>.
+<LI> <A NAME="HodasMiller94"></A>
+     [Hodas and Miller, 1994]<BR> 
+     <A HREF="http://www.cs.hmc.edu/~hodas/">J. Hodas</A>,
+     <A HREF="http://www.cse.psu.edu/~dale/">D. Miller</A>,
+     Logic Programming in a Fragment of Intuitionistic Linear Logic, 
+     Information and Computation 110(2), pp 327-365.
+     <A HREF="ftp://ftp.cse.psu.edu/pub/dale/ic94.ps.Z">Available electronically</a>.
+<LI> <A NAME="OderskyWadler97"></A>
+     [Odersky and Wadler, 1997]<BR> 
+     <A HREF="http://www.cis.unisa.edu.au/~cismxo">M. Odersky</A>,
+     <A HREF="http://cm.bell-labs.com/cm/cs/who/wadler/">P. Wadler</A>,
+     Pizza into Java: Translating Theory into Practice, 
+     In Proceedings of the 24th ACM Symposium on Principles of Programming Languages.
+     <A HREF="http://www.cis.unisa.edu.au/~cismxo/papers/popl97.dvi.gz">Available electronically</a>.
+<LI> <A NAME="TroelstraSchwichtenberg96"></A>
+     [Troelstra and Schwichtenberg, 1996]<BR>
+     <A HREF="http://turing.wins.uva.nl/~anne/">A. Troelstra</A>, 
+     <A HREF="http://www.mathematik.uni-muenchen.de/~gadmin6/professoren/schwichtenberg">H. Schwichtenberg</A>,  
+     Basic Proof Theory, Cambridge Tracts in Theoretical Computer Science,
+     Cambridge University Press.
+</UL>
+
+
+<HR>
+<ADDRESS>
+<A HREF="mailto:Christian.Urban@cl.cam.ac.uk">Christian Urban</A></ADDRESS>
+
+
+<P><!-- Created: Tue Mar  4 00:23:25 GMT 1997 -->
+<!-- hhmts start -->
+Last modified: Sun Sep 23 12:04:47 BST 2001
+<!-- hhmts end -->
+</BODY>
+</HTML>