src/FOL/ex/Prolog.thy
author wenzelm
Sat, 01 Jul 2000 19:55:22 +0200
changeset 9230 17ae63f82ad8
parent 1473 e8d4606e6502
child 17245 1c519a3cca59
permissions -rw-r--r--
GPLed;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1473
e8d4606e6502 expanded tabs
clasohm
parents: 1322
diff changeset
     1
(*  Title:      FOL/ex/prolog.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1473
e8d4606e6502 expanded tabs
clasohm
parents: 1322
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
First-Order Logic: PROLOG examples
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
Inherits from FOL the class term, the type o, and the coercion Trueprop
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
Prolog = FOL +
352
fd3ab8bcb69d removal of obsolete type-declaration syntax
lcp
parents: 0
diff changeset
    12
types   'a list
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
arities list    :: (term)term
1322
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 352
diff changeset
    14
consts  Nil     :: 'a list
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 352
diff changeset
    15
        ":"     :: ['a, 'a list]=> 'a list              (infixr 60)
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 352
diff changeset
    16
        app     :: ['a list, 'a list, 'a list] => o  
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 352
diff changeset
    17
        rev     :: ['a list, 'a list] => o  
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
rules   appNil  "app(Nil,ys,ys)"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
        appCons "app(xs,ys,zs) ==> app(x:xs, ys, x:zs)"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
        revNil  "rev(Nil,Nil)"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
        revCons "[| rev(xs,ys);  app(ys, x:Nil, zs) |] ==> rev(x:xs, zs)"
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
end