author | paulson |
Tue, 11 May 2004 10:47:15 +0200 | |
changeset 14732 | 967db86e853c |
parent 13208 | 965f95a3abd9 |
child 14981 | e73f8140af78 |
permissions | -rw-r--r-- |
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 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
|
5 |
||
6 |
basic examples |
|
7 |
*) |
|
9015 | 8 |
|
9 |
Test = HOHH + |
|
10 |
||
11 |
types nat |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
9015
diff
changeset
|
12 |
arities nat :: type |
9015 | 13 |
|
14 |
types 'a list |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
9015
diff
changeset
|
15 |
arities list :: (type) type |
9015 | 16 |
|
17 |
consts Nil :: 'a list ("[]") |
|
18 |
Cons :: 'a => 'a list => 'a list (infixr "#" 65) |
|
19 |
||
20 |
syntax |
|
21 |
(* list Enumeration *) |
|
22 |
"@list" :: args => 'a list ("[(_)]") |
|
23 |
||
24 |
translations |
|
25 |
"[x, xs]" == "x#[xs]" |
|
26 |
"[x]" == "x#[]" |
|
27 |
||
28 |
types person |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
9015
diff
changeset
|
29 |
arities person :: type |
9015 | 30 |
|
31 |
consts |
|
32 |
append :: ['a list, 'a list, 'a list] => bool |
|
33 |
reverse :: ['a list, 'a list] => bool |
|
34 |
||
35 |
mappred :: [('a => 'b => bool), 'a list, 'b list] => bool |
|
36 |
mapfun :: [('a => 'b), 'a list, 'b list] => bool |
|
37 |
||
38 |
bob :: person |
|
39 |
sue :: person |
|
40 |
ned :: person |
|
41 |
||
42 |
"23" :: nat ("23") |
|
43 |
"24" :: nat ("24") |
|
44 |
"25" :: nat ("25") |
|
45 |
||
46 |
age :: [person, nat] => bool |
|
47 |
||
48 |
eq :: ['a, 'a] => bool |
|
49 |
||
50 |
empty :: ['b] => bool |
|
51 |
add :: ['a, 'b, 'b] => bool |
|
52 |
remove :: ['a, 'b, 'b] => bool |
|
53 |
bag_appl:: ['a, 'a, 'a, 'a] => bool |
|
54 |
||
55 |
rules 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 [])" |
|
61 |
||
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" |
|
66 |
||
67 |
age "age bob 24 .. |
|
68 |
age sue 23 .. |
|
69 |
age ned 23" |
|
70 |
||
71 |
eq "eq x x" |
|
72 |
||
73 |
(* actual definitions of empty and add is hidden -> yields abstract data type *) |
|
74 |
||
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 |
end |