author | haftmann |
Thu, 16 Jan 2025 18:07:31 +0100 | |
changeset 81818 | 1085eb118dc7 |
parent 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
13208 | 1 |
(* Title: HOL/Prolog/Func.thy |
2 |
Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur) |
|
3 |
*) |
|
9015 | 4 |
|
63167 | 5 |
section \<open>Untyped functional language, with call by value semantics\<close> |
9015 | 6 |
|
17311 | 7 |
theory Func |
35301
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
haftmann
parents:
35265
diff
changeset
|
8 |
imports HOHH |
17311 | 9 |
begin |
9015 | 10 |
|
17311 | 11 |
typedecl tm |
9015 | 12 |
|
35265 | 13 |
axiomatization |
14 |
abs :: "(tm \<Rightarrow> tm) \<Rightarrow> tm" and |
|
15 |
app :: "tm \<Rightarrow> tm \<Rightarrow> tm" and |
|
17311 | 16 |
|
35265 | 17 |
cond :: "tm \<Rightarrow> tm \<Rightarrow> tm \<Rightarrow> tm" and |
18 |
"fix" :: "(tm \<Rightarrow> tm) \<Rightarrow> tm" and |
|
9015 | 19 |
|
35265 | 20 |
true :: tm and |
21 |
false :: tm and |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
63167
diff
changeset
|
22 |
"and" :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixr \<open>and\<close> 999) and |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
63167
diff
changeset
|
23 |
eq :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixr \<open>eq\<close> 999) and |
35265 | 24 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
63167
diff
changeset
|
25 |
Z :: tm (\<open>Z\<close>) and |
35265 | 26 |
S :: "tm \<Rightarrow> tm" and |
9015 | 27 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
63167
diff
changeset
|
28 |
plus :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixl \<open>+\<close> 65) and |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
63167
diff
changeset
|
29 |
minus :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixl \<open>-\<close> 65) and |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
63167
diff
changeset
|
30 |
times :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixl \<open>*\<close> 70) and |
9015 | 31 |
|
35265 | 32 |
eval :: "tm \<Rightarrow> tm \<Rightarrow> bool" where |
33 |
||
34 |
eval: " |
|
9015 | 35 |
|
36 |
eval (abs RR) (abs RR).. |
|
37 |
eval (app F X) V :- eval F (abs R) & eval X U & eval (R U) V.. |
|
38 |
||
39 |
eval (cond P L1 R1) D1 :- eval P true & eval L1 D1.. |
|
40 |
eval (cond P L2 R2) D2 :- eval P false & eval R2 D2.. |
|
41 |
eval (fix G) W :- eval (G (fix G)) W.. |
|
42 |
||
43 |
eval true true .. |
|
44 |
eval false false.. |
|
45 |
eval (P and Q) true :- eval P true & eval Q true .. |
|
46 |
eval (P and Q) false :- eval P false | eval Q false.. |
|
17311 | 47 |
eval (A1 eq B1) true :- eval A1 C1 & eval B1 C1.. |
9015 | 48 |
eval (A2 eq B2) false :- True.. |
49 |
||
50 |
eval Z Z.. |
|
51 |
eval (S N) (S M) :- eval N M.. |
|
52 |
eval ( Z + M) K :- eval M K.. |
|
53 |
eval ((S N) + M) (S K) :- eval (N + M) K.. |
|
54 |
eval (N - Z) K :- eval N K.. |
|
55 |
eval ((S N) - (S M)) K :- eval (N- M) K.. |
|
56 |
eval ( Z * M) Z.. |
|
57 |
eval ((S N) * M) K :- eval (N * M) L & eval (L + M) K" |
|
58 |
||
21425 | 59 |
lemmas prog_Func = eval |
60 |
||
61337 | 61 |
schematic_goal "eval ((S (S Z)) + (S Z)) ?X" |
21425 | 62 |
apply (prolog prog_Func) |
63 |
done |
|
64 |
||
61337 | 65 |
schematic_goal "eval (app (fix (%fact. abs(%n. cond (n eq Z) (S Z) |
21425 | 66 |
(n * (app fact (n - (S Z))))))) (S (S (S Z)))) ?X" |
67 |
apply (prolog prog_Func) |
|
68 |
done |
|
17311 | 69 |
|
9015 | 70 |
end |