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