Theory PCF
section ‹Logical relations for definability in PCF›
theory PCF
imports
Basis
Logical_Relations
begin
text‹
\label{sec:directsem}
Using this machinery we can demonstrate some classical results about
PCF \<^citep>‹"Plotkin77"›. We diverge from the traditional treatment by
considering PCF as an untyped language and including both call-by-name
(CBN) and call-by-value (CBV) abstractions following
\<^citet>‹"DBLP:conf/icalp/Reynolds74"›. We also adopt some of the
presentation of \<^citet>‹‹Chapter~11› in "Winskel:1993"›, in particular by
making the fixed point operator a binding construct.
We model the syntax of PCF as a HOL datatype, where variables have
names drawn from the naturals:
›
type_synonym var = nat