author | wenzelm |
Wed, 13 Mar 2019 13:46:16 +0100 | |
changeset 69904 | 6f5bd59f75f4 |
parent 69605 | a96320074298 |
child 76091 | 922e3f9251ac |
permissions | -rw-r--r-- |
13208 | 1 |
(* Title: HOL/Prolog/HOHH.thy |
2 |
Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur) |
|
3 |
*) |
|
9015 | 4 |
|
63167 | 5 |
section \<open>Higher-order hereditary Harrop formulas\<close> |
17311 | 6 |
|
7 |
theory HOHH |
|
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63167
diff
changeset
|
8 |
imports HOL.HOL |
17311 | 9 |
begin |
9015 | 10 |
|
69605 | 11 |
ML_file \<open>prolog.ML\<close> |
48891 | 12 |
|
21425 | 13 |
method_setup ptac = |
63167 | 14 |
\<open>Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (Prolog.ptac ctxt thms))\<close> |
42814 | 15 |
"basic Lambda Prolog interpreter" |
21425 | 16 |
|
17 |
method_setup prolog = |
|
63167 | 18 |
\<open>Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD (Prolog.prolog_tac ctxt thms))\<close> |
21425 | 19 |
"Lambda Prolog interpreter" |
20 |
||
9015 | 21 |
consts |
22 |
||
17311 | 23 |
(* D-formulas (programs): D ::= !x. D | D .. D | D :- G | A *) |
21425 | 24 |
Dand :: "[bool, bool] => bool" (infixr ".." 28) |
25 |
Dif :: "[bool, bool] => bool" (infixl ":-" 29) |
|
9015 | 26 |
|
17311 | 27 |
(* G-formulas (goals): G ::= A | G & G | G | G | ? x. G |
28 |
| True | !x. G | D => G *) |
|
21425 | 29 |
(*Dand' :: "[bool, bool] => bool" (infixr "," 35)*) |
30 |
Dimp :: "[bool, bool] => bool" (infixr "=>" 27) |
|
9015 | 31 |
|
32 |
translations |
|
33 |
||
17311 | 34 |
"D :- G" => "G --> D" |
35 |
"D1 .. D2" => "D1 & D2" |
|
36 |
(*"G1 , G2" => "G1 & G2"*) |
|
37 |
"D => G" => "D --> G" |
|
9015 | 38 |
|
39 |
end |