Theory Propositional_Logic_Class.Classical_Logic_Completeness
chapter ‹ Classical Soundness and Completeness \label{sec:classical-propositional-calculus} ›
theory Classical_Logic_Completeness
imports Classical_Logic
begin
text ‹ The following presents soundness completeness of the classical propositional
calculus for propositional semantics. The classical propositional calculus
is sometimes referred to as the ∗‹sentential calculus›.
We give a concrete algebraic data type for propositional
formulae in \S\ref{sec:classical-calculus-syntax}. We inductively
define a logical judgement ‹⊢⇩p⇩r⇩o⇩p› for these formulae.
We also define the Tarski truth relation ‹⊨⇩p⇩r⇩o⇩p› inductively,
which we present in \S\ref{sec:propositional-semantics}.›
text ‹ The most significant results here are the ∗‹embedding theorems›.
These theorems show that the propositional calculus
can be embedded in any logic extending @{class classical_logic}.
These theorems are proved in \S\ref{sec:propositional-embedding}. ›
section ‹ Syntax \label{sec:classical-calculus-syntax} ›
text ‹ Here we provide the usual language for formulae in the propositional
calculus. It contains ∗‹falsum› ‹❙⊥›, implication ‹(❙→)›, and a way of
constructing ∗‹atomic› propositions ‹λ φ . ❙⟨ φ ❙⟩›. Defining the
language is straight-forward using an algebraic data type. ›