| author | paulson | 
| Mon, 15 Mar 2004 10:58:49 +0100 | |
| changeset 14470 | 1ffe42cfaefe | 
| parent 13208 | 965f95a3abd9 | 
| child 14981 | e73f8140af78 | 
| permissions | -rw-r--r-- | 
| 13208 | 1 | (* Title: HOL/Prolog/Func.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 | untyped functional language, with call by value semantics | |
| 7 | *) | |
| 9015 | 8 | |
| 9 | Func = HOHH + | |
| 10 | ||
| 11 | types tm | |
| 12 | ||
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
9015diff
changeset | 13 | arities tm :: type | 
| 9015 | 14 | |
| 15 | consts abs :: (tm => tm) => tm | |
| 16 | app :: tm => tm => tm | |
| 17 | ||
| 18 | cond :: tm => tm => tm => tm | |
| 19 | fix :: (tm => tm) => tm | |
| 20 | ||
| 21 | true, | |
| 22 | false :: tm | |
| 23 | "and" :: tm => tm => tm (infixr 999) | |
| 24 | "eq" :: tm => tm => tm (infixr 999) | |
| 25 | ||
| 26 | 	"0"	:: tm			("Z")
 | |
| 27 | S :: tm => tm | |
| 28 | (* | |
| 29 | "++", "--", | |
| 30 | "**" :: tm => tm => tm (infixr 999) | |
| 31 | *) | |
| 32 | eval :: [tm, tm] => bool | |
| 33 | ||
| 34 | arities tm :: plus | |
| 35 | arities tm :: minus | |
| 36 | arities tm :: times | |
| 37 | ||
| 38 | rules eval " | |
| 39 | ||
| 40 | eval (abs RR) (abs RR).. | |
| 41 | eval (app F X) V :- eval F (abs R) & eval X U & eval (R U) V.. | |
| 42 | ||
| 43 | eval (cond P L1 R1) D1 :- eval P true & eval L1 D1.. | |
| 44 | eval (cond P L2 R2) D2 :- eval P false & eval R2 D2.. | |
| 45 | eval (fix G) W :- eval (G (fix G)) W.. | |
| 46 | ||
| 47 | eval true true .. | |
| 48 | eval false false.. | |
| 49 | eval (P and Q) true :- eval P true & eval Q true .. | |
| 50 | eval (P and Q) false :- eval P false | eval Q false.. | |
| 51 | eval (A1 eq B1) true :- eval A1 C1 & eval B1 C1.. | |
| 52 | eval (A2 eq B2) false :- True.. | |
| 53 | ||
| 54 | eval Z Z.. | |
| 55 | eval (S N) (S M) :- eval N M.. | |
| 56 | eval ( Z + M) K :- eval M K.. | |
| 57 | eval ((S N) + M) (S K) :- eval (N + M) K.. | |
| 58 | eval (N - Z) K :- eval N K.. | |
| 59 | eval ((S N) - (S M)) K :- eval (N- M) K.. | |
| 60 | eval ( Z * M) Z.. | |
| 61 | eval ((S N) * M) K :- eval (N * M) L & eval (L + M) K" | |
| 62 | ||
| 63 | end |