| 0 |      1 | (*  Title: 	FOL/ex/prolog.thy
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
|  |      4 |     Copyright   1992  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | First-Order Logic: PROLOG examples
 | 
|  |      7 | 
 | 
|  |      8 | Inherits from FOL the class term, the type o, and the coercion Trueprop
 | 
|  |      9 | *)
 | 
|  |     10 | 
 | 
|  |     11 | Prolog = FOL +
 | 
|  |     12 | types   list 1
 | 
|  |     13 | arities list    :: (term)term
 | 
|  |     14 | consts  Nil     :: "'a list"
 | 
|  |     15 |         ":"     :: "['a, 'a list]=> 'a list"            (infixr 60)
 | 
|  |     16 |         app     :: "['a list, 'a list, 'a list] => o"
 | 
|  |     17 |         rev     :: "['a list, 'a list] => o"
 | 
|  |     18 | rules   appNil  "app(Nil,ys,ys)"
 | 
|  |     19 |         appCons "app(xs,ys,zs) ==> app(x:xs, ys, x:zs)"
 | 
|  |     20 |         revNil  "rev(Nil,Nil)"
 | 
|  |     21 |         revCons "[| rev(xs,ys);  app(ys, x:Nil, zs) |] ==> rev(x:xs, zs)"
 | 
|  |     22 | end
 |