src/FOL/ex/Prolog.thy
author wenzelm
Sat, 15 Oct 2005 00:08:00 +0200
changeset 17851 2fa4f9b54761
parent 17245 1c519a3cca59
child 19819 14de4d05d275
permissions -rw-r--r--
tuned comments;
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
17245
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
     7
header {* First-Order Logic: PROLOG examples *}
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
     8
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
     9
theory Prolog
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    10
imports FOL
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    11
begin
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    12
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    13
typedecl 'a list
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    14
arities list :: ("term") "term"
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    15
consts
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    16
  Nil     :: "'a list"
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    17
  Cons    :: "['a, 'a list]=> 'a list"    (infixr ":" 60)
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    18
  app     :: "['a list, 'a list, 'a list] => o"
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    19
  rev     :: "['a list, 'a list] => o"
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    20
axioms
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    21
  appNil:  "app(Nil,ys,ys)"
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    22
  appCons: "app(xs,ys,zs) ==> app(x:xs, ys, x:zs)"
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    23
  revNil:  "rev(Nil,Nil)"
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    24
  revCons: "[| rev(xs,ys);  app(ys, x:Nil, zs) |] ==> rev(x:xs, zs)"
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    25
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    26
ML {* use_legacy_bindings (the_context ()) *}
1c519a3cca59 converted to Isar theory format;
wenzelm
parents: 1473
diff changeset
    27
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
end