src/HOL/Prolog/Test.thy
author wenzelm
Wed, 07 Sep 2005 21:07:09 +0200
changeset 17311 5b1d47d920ce
parent 14981 e73f8140af78
child 21425 c11ab38b78a7
permissions -rw-r--r--
converted to Isar theory format;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13208
965f95a3abd9 added the usual file headers
oheimb
parents: 12338
diff changeset
     1
(*  Title:    HOL/Prolog/Test.thy
965f95a3abd9 added the usual file headers
oheimb
parents: 12338
diff changeset
     2
    ID:       $Id$
965f95a3abd9 added the usual file headers
oheimb
parents: 12338
diff changeset
     3
    Author:   David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
965f95a3abd9 added the usual file headers
oheimb
parents: 12338
diff changeset
     4
*)
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
     5
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     6
header {* Basic examples *}
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
     7
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     8
theory Test
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
     9
imports HOHH
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    10
begin
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    11
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    12
typedecl nat
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    13
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    14
typedecl 'a list
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    15
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    16
consts
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    17
  Nil   :: "'a list"                                  ("[]")
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    18
  Cons  :: "'a => 'a list => 'a list"                 (infixr "#"  65)
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    19
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    20
syntax
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    21
  (* list Enumeration *)
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    22
  "@list"     :: "args => 'a list"                          ("[(_)]")
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    23
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    24
translations
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    25
  "[x, xs]"     == "x#[xs]"
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    26
  "[x]"         == "x#[]"
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    27
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    28
typedecl person
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    29
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    30
consts
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    31
  append  :: "['a list, 'a list, 'a list]            => bool"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    32
  reverse :: "['a list, 'a list]                     => bool"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    33
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    34
  mappred :: "[('a => 'b => bool), 'a list, 'b list] => bool"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    35
  mapfun  :: "[('a => 'b), 'a list, 'b list]         => bool"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    36
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    37
  bob     :: person
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    38
  sue     :: person
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    39
  ned     :: person
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    40
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    41
  "23"    :: nat          ("23")
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    42
  "24"    :: nat          ("24")
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    43
  "25"    :: nat          ("25")
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    44
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    45
  age     :: "[person, nat]                          => bool"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    46
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    47
  eq      :: "['a, 'a]                               => bool"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    48
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    49
  empty   :: "['b]                                   => bool"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    50
  add     :: "['a, 'b, 'b]                           => bool"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    51
  remove  :: "['a, 'b, 'b]                           => bool"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    52
  bag_appl:: "['a, 'a, 'a, 'a]                       => bool"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    53
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    54
axioms
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    55
        append:  "append  []    xs  xs    ..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    56
                  append (x#xs) ys (x#zs) :- append xs ys zs"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    57
        reverse: "reverse L1 L2 :- (!rev_aux.
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    58
                  (!L.          rev_aux  []    L  L )..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    59
                  (!X L1 L2 L3. rev_aux (X#L1) L2 L3 :- rev_aux L1 L2 (X#L3))
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    60
                  => rev_aux L1 L2 [])"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    61
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    62
        mappred: "mappred P  []     []    ..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    63
                  mappred P (x#xs) (y#ys) :- P x y  &  mappred P xs ys"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    64
        mapfun:  "mapfun f  []     []      ..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    65
                  mapfun f (x#xs) (f x#ys) :- mapfun f xs ys"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    66
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    67
        age:     "age bob 24 ..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    68
                  age sue 23 ..
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    69
                  age ned 23"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    70
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    71
        eq:      "eq x x"
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    72
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    73
(* actual definitions of empty and add is hidden -> yields abstract data type *)
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    74
17311
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    75
        bag_appl: "bag_appl A B X Y:- (? S1 S2 S3 S4 S5.
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    76
                                empty    S1    &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    77
                                add A    S1 S2 &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    78
                                add B    S2 S3 &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    79
                                remove X S3 S4 &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    80
                                remove Y S4 S5 &
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    81
                                empty    S5)"
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    82
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    83
ML {* use_legacy_bindings (the_context ()) *}
5b1d47d920ce converted to Isar theory format;
wenzelm
parents: 14981
diff changeset
    84
9015
8006e9009621 added HOL/Prolog
oheimb
parents:
diff changeset
    85
end