13208
|
1 |
(* Title: HOL/Prolog/HOHH.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur)
|
|
4 |
*)
|
9015
|
5 |
|
17311
|
6 |
header {* Higher-order hereditary Harrop formulas *}
|
|
7 |
|
|
8 |
theory HOHH
|
|
9 |
imports HOL
|
|
10 |
begin
|
9015
|
11 |
|
|
12 |
consts
|
|
13 |
|
17311
|
14 |
(* D-formulas (programs): D ::= !x. D | D .. D | D :- G | A *)
|
|
15 |
"Dand" :: "[bool, bool] => bool" (infixr ".." 28)
|
|
16 |
":-" :: "[bool, bool] => bool" (infixl 29)
|
9015
|
17 |
|
17311
|
18 |
(* G-formulas (goals): G ::= A | G & G | G | G | ? x. G
|
|
19 |
| True | !x. G | D => G *)
|
|
20 |
(*"," :: "[bool, bool] => bool" (infixr 35)*)
|
|
21 |
"=>" :: "[bool, bool] => bool" (infixr 27)
|
9015
|
22 |
|
|
23 |
translations
|
|
24 |
|
17311
|
25 |
"D :- G" => "G --> D"
|
|
26 |
"D1 .. D2" => "D1 & D2"
|
|
27 |
(*"G1 , G2" => "G1 & G2"*)
|
|
28 |
"D => G" => "D --> G"
|
9015
|
29 |
|
|
30 |
end
|