Theory Logic
theory Logic
imports LaTeXsugar
begin
text‹
\vspace{-5ex}
\section{Formulas}
The core syntax of formulas (\textit{form} below)
provides the standard logical constructs, in decreasing order of precedence:
\[
\begin{array}{rcl}
\mathit{form} & ::= &
‹(form)› ~\mid~
\<^const>‹True› ~\mid~
\<^const>‹False› ~\mid~
\<^prop>‹term = term›\\
&\mid& \<^prop>‹¬ form›\index{$HOL4@\isasymnot} ~\mid~
\<^prop>‹form ∧ form›\index{$HOL0@\isasymand} ~\mid~
\<^prop>‹form ∨ form›\index{$HOL1@\isasymor} ~\mid~
\<^prop>‹form ⟶ form›\index{$HOL2@\isasymlongrightarrow}\\
&\mid& \<^prop>‹∀x. form›\index{$HOL6@\isasymforall} ~\mid~ \<^prop>‹∃x. form›\index{$HOL7@\isasymexists}
\end{array}
\]
Terms are the ones we have seen all along, built from constants, variables,
function application and ‹λ›-abstraction, including all the syntactic
sugar like infix symbols, ‹if›, ‹case›, etc.
\begin{warn}
Remember that formulas are simply terms of type ‹bool›. Hence
‹=› also works for formulas. Beware that ‹=› has a higher
precedence than the other logical operators. Hence \<^prop>‹s = t ∧ A› means
‹(s = t) ∧ A›, and \<^prop>‹A∧B = B∧A› means ‹A ∧ (B = B) ∧ A›.
Logical equivalence can also be written with
‹⟷› instead of ‹=›, where ‹⟷› has the same low
precedence as ‹⟶›. Hence ‹A ∧ B ⟷ B ∧ A› really means
‹(A ∧ B) ⟷ (B ∧ A)›.
\end{warn}
\begin{warn}
Quantifiers need to be enclosed in parentheses if they are nested within
other constructs (just like ‹if›, ‹case› and ‹let›).
\end{warn}
The most frequent logical symbols and their ASCII representations are shown
in Fig.~\ref{fig:log-symbols}.
\begin{figure}
\begin{center}
\begin{tabular}{l@ {\qquad}l@ {\qquad}l}
‹∀› & \xsymbol{forall} & \texttt{ALL}\\
‹∃› & \xsymbol{exists} & \texttt{EX}\\
‹λ› & \xsymbol{lambda} & \texttt{\%}\\
‹⟶› & \texttt{-{\kern0pt}->}\\
‹⟷› & \texttt{<->}\\
‹∧› & \texttt{/\char`\\} & \texttt{\&}\\
‹∨› & \texttt{\char`\\/} & \texttt{|}\\
‹¬› & \xsymbol{not} & \texttt{\char`~}\\
‹≠› & \xsymbol{noteq} & \texttt{\char`~=}
\end{tabular}
\end{center}
\caption{Logical symbols and their ASCII forms}
\label{fig:log-symbols}
\end{figure}
The first column shows the symbols, the other columns ASCII representations.
The \texttt{\char`\\}\texttt{<...>} form is always converted into the symbolic form
by the Isabelle interfaces, the treatment of the other ASCII forms
depends on the interface. The ASCII forms \texttt{/\char`\\} and
\texttt{\char`\\/}
are special in that they are merely keyboard shortcuts for the interface and
not logical symbols by themselves.
\begin{warn}
The implication ‹⟹› is part of the Isabelle framework. It structures
theorems and proof states, separating assumptions from conclusions.
The implication ‹⟶› is part of the logic HOL and can occur inside the
formulas that make up the assumptions and conclusion.
Theorems should be of the form ‹⟦ A⇩1; …; A⇩n ⟧ ⟹ A›,
not ‹A⇩1 ∧ … ∧ A⇩n ⟶ A›. Both are logically equivalent
but the first one works better when using the theorem in further proofs.
The ASCII representation of ‹⟦› and ‹⟧› is \texttt{[|} and \texttt{|]}.
\end{warn}
\section{Sets}
\label{sec:Sets}
Sets of elements of type \<^typ>‹'a› have type \<^typ>‹'a set›\index{set@‹set›}.
They can be finite or infinite. Sets come with the usual notation:
\begin{itemize}
\item \indexed{\<^term>‹{}›}{$IMP042},\quad ‹{e⇩1,…,e⇩n}›
\item \<^prop>‹e ∈ A›\index{$HOLSet0@\isasymin},\quad \<^prop>‹A ⊆ B›\index{$HOLSet2@\isasymsubseteq}
\item \<^term>‹A ∪ B›\index{$HOLSet4@\isasymunion},\quad \<^term>‹A ∩ B›\index{$HOLSet5@\isasyminter},\quad \<^term>‹A - B›,\quad \<^term>‹-A›
\end{itemize}
(where \<^term>‹A-B› and ‹-A› are set difference and complement)
and much more. \<^const>‹UNIV› is the set of all elements of some type.
Set comprehension\index{set comprehension} is written
\<^term>‹{x. P}›\index{$IMP042@\<^term>‹{x. P}›} rather than ‹{x | P}›.
\begin{warn}
In \<^term>‹{x. P}› the ‹x› must be a variable. Set comprehension
involving a proper term ‹t› must be written
\noquotes{@{term[source] "{t | x y. P}"}}\index{$IMP042@‹{t |x. P}›},
where ‹x y› are those free variables in ‹t›
that occur in ‹P›.
This is just a shorthand for \<^term>‹{v. ∃x y. v = t ∧ P}›, where
‹v› is a new variable. For example, \<^term>‹{x+y|x. x ∈ A}›
is short for \noquotes{@{term[source]"{v. ∃x. v = x+y ∧ x ∈ A}"}}.
\end{warn}
Here are the ASCII representations of the mathematical symbols:
\begin{center}
\begin{tabular}{l@ {\quad}l@ {\quad}l}
‹∈› & \texttt{\char`\\\char`∈} & \texttt{:}\\
‹⊆› & \texttt{\char`\\\char`⊆} & \texttt{<=}\\
‹∪› & \texttt{\char`\\\char`∪} & \texttt{Un}\\
‹∩› & \texttt{\char`\\\char`∩} & \texttt{Int}
\end{tabular}
\end{center}
Sets also allow bounded quantifications \<^prop>‹∀x ∈ A. P› and
\<^prop>‹∃x ∈ A. P›.
For the more ambitious, there are also ‹⋃›\index{$HOLSet6@\isasymUnion}
and ‹⋂›\index{$HOLSet7@\isasymInter}:
\begin{center}
@{thm Union_eq} \qquad @{thm Inter_eq}
\end{center}
The ASCII forms of ‹⋃› are \texttt{\char`\\\char`⋃} and \texttt{Union},
those of ‹⋂› are \texttt{\char`\\\char`⋂} and \texttt{Inter}.
There are also indexed unions and intersections:
\begin{center}
@{thm[eta_contract=false] UNION_eq} \\ @{thm[eta_contract=false] INTER_eq}
\end{center}
The ASCII forms are \ \texttt{UN x:A.~B} \ and \ \texttt{INT x:A. B} \
where \texttt{x} may occur in \texttt{B}.
If \texttt{A} is \texttt{UNIV} you can write \ \texttt{UN x.~B} \ and \ \texttt{INT x. B}.
Some other frequently useful functions on sets are the following:
\begin{center}
\begin{tabular}{l@ {\quad}l}
@{const_typ set}\index{set@\<^const>‹set›} & converts a list to the set of its elements\\
@{const_typ finite}\index{finite@\<^const>‹finite›} & is true iff its argument is finite\\
\noquotes{@{term[source] "card :: 'a set ⇒ nat"}}\index{card@\<^const>‹card›} & is the cardinality of a finite set\\
& and is ‹0› for all infinite sets\\
@{thm image_def}\index{$IMP042@\<^term>‹f ` A›} & is the image of a function over a set
\end{tabular}
\end{center}
See \<^cite>‹"Nipkow-Main"› for the wealth of further predefined functions in theory
\<^theory>‹Main›.
\subsection*{Exercises}
\exercise
Start from the data type of binary trees defined earlier:
›