| 47269 |      1 | (*<*)
 | 
|  |      2 | theory Basics
 | 
|  |      3 | imports Main
 | 
|  |      4 | begin
 | 
|  |      5 | (*>*)
 | 
|  |      6 | text{*
 | 
|  |      7 | This chapter introduces HOL as a functional programming language and shows
 | 
|  |      8 | how to prove properties of functional programs by induction.
 | 
|  |      9 | 
 | 
|  |     10 | \section{Basics}
 | 
|  |     11 | 
 | 
|  |     12 | \subsection{Types, Terms and Formulae}
 | 
|  |     13 | \label{sec:TypesTermsForms}
 | 
|  |     14 | 
 | 
|  |     15 | HOL is a typed logic whose type system resembles that of functional
 | 
|  |     16 | programming languages. Thus there are
 | 
|  |     17 | \begin{description}
 | 
|  |     18 | \item[base types,] 
 | 
|  |     19 | in particular @{typ bool}, the type of truth values,
 | 
|  |     20 | @{typ nat}, the type of natural numbers ($\mathbb{N}$), and @{typ int},
 | 
|  |     21 | the type of mathematical integers ($\mathbb{Z}$).
 | 
|  |     22 | \item[type constructors,]
 | 
|  |     23 |  in particular @{text list}, the type of
 | 
|  |     24 | lists, and @{text set}, the type of sets. Type constructors are written
 | 
|  |     25 | postfix, e.g.\ @{typ "nat list"} is the type of lists whose elements are
 | 
|  |     26 | natural numbers.
 | 
|  |     27 | \item[function types,]
 | 
|  |     28 | denoted by @{text"\<Rightarrow>"}.
 | 
|  |     29 | \item[type variables,]
 | 
|  |     30 |   denoted by @{typ 'a}, @{typ 'b} etc., just like in ML\@.
 | 
|  |     31 | \end{description}
 | 
|  |     32 | 
 | 
|  |     33 | \concept{Terms} are formed as in functional programming by
 | 
|  |     34 | applying functions to arguments. If @{text f} is a function of type
 | 
|  |     35 | @{text"\<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2"} and @{text t} is a term of type
 | 
|  |     36 | @{text"\<tau>\<^isub>1"} then @{term"f t"} is a term of type @{text"\<tau>\<^isub>2"}. We write @{text "t :: \<tau>"} to mean that term @{text t} has type @{text \<tau>}.
 | 
|  |     37 | 
 | 
|  |     38 | \begin{warn}
 | 
|  |     39 | There are many predefined infix symbols like @{text "+"} and @{text"\<le>"}.
 | 
|  |     40 | The name of the corresponding binary function is @{term"op +"},
 | 
|  |     41 | not just @{text"+"}. That is, @{term"x + y"} is syntactic sugar for
 | 
|  |     42 | \noquotes{@{term[source]"op + x y"}}.
 | 
|  |     43 | \end{warn}
 | 
|  |     44 | 
 | 
|  |     45 | HOL also supports some basic constructs from functional programming:
 | 
|  |     46 | \begin{quote}
 | 
|  |     47 | @{text "(if b then t\<^isub>1 else t\<^isub>2)"}\\
 | 
|  |     48 | @{text "(let x = t in u)"}\\
 | 
|  |     49 | @{text "(case t of pat\<^isub>1 \<Rightarrow> t\<^isub>1 | \<dots> | pat\<^isub>n \<Rightarrow> t\<^isub>n)"}
 | 
|  |     50 | \end{quote}
 | 
|  |     51 | \begin{warn}
 | 
|  |     52 | The above three constructs must always be enclosed in parentheses
 | 
|  |     53 | if they occur inside other constructs.
 | 
|  |     54 | \end{warn}
 | 
|  |     55 | Terms may also contain @{text "\<lambda>"}-abstractions. For example,
 | 
|  |     56 | @{term "\<lambda>x. x"} is the identity function.
 | 
|  |     57 | 
 | 
|  |     58 | \concept{Formulae} are terms of type @{text bool}.
 | 
|  |     59 | There are the basic constants @{term True} and @{term False} and
 | 
|  |     60 | the usual logical connectives (in decreasing order of precedence):
 | 
|  |     61 | @{text"\<not>"}, @{text"\<and>"}, @{text"\<or>"}, @{text"\<longrightarrow>"}.
 | 
|  |     62 | 
 | 
|  |     63 | \concept{Equality} is available in the form of the infix function @{text "="}
 | 
|  |     64 | of type @{typ "'a \<Rightarrow> 'a \<Rightarrow> bool"}. It also works for formulas, where
 | 
|  |     65 | it means ``if and only if''.
 | 
|  |     66 | 
 | 
|  |     67 | \concept{Quantifiers} are written @{prop"\<forall>x. P"} and @{prop"\<exists>x. P"}.
 | 
|  |     68 | 
 | 
|  |     69 | Isabelle automatically computes the type of each variable in a term. This is
 | 
|  |     70 | called \concept{type inference}.  Despite type inference, it is sometimes
 | 
|  |     71 | necessary to attach explicit \concept{type constraints} (or \concept{type
 | 
|  |     72 | annotations}) to a variable or term.  The syntax is @{text "t :: \<tau>"} as in
 | 
|  |     73 | \mbox{\noquotes{@{prop[source] "m < (n::nat)"}}}. Type constraints may be
 | 
|  |     74 | needed to
 | 
|  |     75 | disambiguate terms involving overloaded functions such as @{text "+"}, @{text
 | 
|  |     76 | "*"} and @{text"\<le>"}.
 | 
|  |     77 | 
 | 
|  |     78 | Finally there are the universal quantifier @{text"\<And>"} and the implication
 | 
|  |     79 | @{text"\<Longrightarrow>"}. They are part of the Isabelle framework, not the logic
 | 
|  |     80 | HOL. Logically, they agree with their HOL counterparts @{text"\<forall>"} and
 | 
|  |     81 | @{text"\<longrightarrow>"}, but operationally they behave differently. This will become
 | 
|  |     82 | clearer as we go along.
 | 
|  |     83 | \begin{warn}
 | 
|  |     84 | Right-arrows of all kinds always associate to the right. In particular,
 | 
|  |     85 | the formula
 | 
|  |     86 | @{text"A\<^isub>1 \<Longrightarrow> A\<^isub>2 \<Longrightarrow> A\<^isub>3"} means @{text "A\<^isub>1 \<Longrightarrow> (A\<^isub>2 \<Longrightarrow> A\<^isub>3)"}.
 | 
|  |     87 | The (Isabelle specific) notation \mbox{@{text"\<lbrakk> A\<^isub>1; \<dots>; A\<^isub>n \<rbrakk> \<Longrightarrow> A"}}
 | 
|  |     88 | is short for the iterated implication \mbox{@{text"A\<^isub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^isub>n \<Longrightarrow> A"}}.
 | 
|  |     89 | Sometimes we also employ inference rule notation:
 | 
|  |     90 | \inferrule{\mbox{@{text "A\<^isub>1"}}\\ \mbox{@{text "\<dots>"}}\\ \mbox{@{text "A\<^isub>n"}}}
 | 
|  |     91 | {\mbox{@{text A}}}
 | 
|  |     92 | \end{warn}
 | 
|  |     93 | 
 | 
|  |     94 | 
 | 
|  |     95 | \subsection{Theories}
 | 
|  |     96 | \label{sec:Basic:Theories}
 | 
|  |     97 | 
 | 
|  |     98 | Roughly speaking, a \concept{theory} is a named collection of types,
 | 
|  |     99 | functions, and theorems, much like a module in a programming language.
 | 
|  |    100 | All the Isabelle text that you ever type needs to go into a theory.
 | 
|  |    101 | The general format of a theory @{text T} is
 | 
|  |    102 | \begin{quote}
 | 
|  |    103 | \isacom{theory} @{text T}\\
 | 
|  |    104 | \isacom{imports} @{text "T\<^isub>1 \<dots> T\<^isub>n"}\\
 | 
|  |    105 | \isacom{begin}\\
 | 
|  |    106 | \emph{definitions, theorems and proofs}\\
 | 
|  |    107 | \isacom{end}
 | 
|  |    108 | \end{quote}
 | 
|  |    109 | where @{text "T\<^isub>1 \<dots> T\<^isub>n"} are the names of existing
 | 
|  |    110 | theories that @{text T} is based on. The @{text "T\<^isub>i"} are the
 | 
|  |    111 | direct \concept{parent theories} of @{text T}.
 | 
|  |    112 | Everything defined in the parent theories (and their parents, recursively) is
 | 
|  |    113 | automatically visible. Each theory @{text T} must
 | 
|  |    114 | reside in a \concept{theory file} named @{text "T.thy"}.
 | 
|  |    115 | 
 | 
|  |    116 | \begin{warn}
 | 
|  |    117 | HOL contains a theory @{text Main}, the union of all the basic
 | 
|  |    118 | predefined theories like arithmetic, lists, sets, etc.
 | 
|  |    119 | Unless you know what you are doing, always include @{text Main}
 | 
|  |    120 | as a direct or indirect parent of all your theories.
 | 
|  |    121 | \end{warn}
 | 
|  |    122 | 
 | 
|  |    123 | In addition to the theories that come with the Isabelle/HOL distribution
 | 
|  |    124 | (see \url{http://isabelle.in.tum.de/library/HOL/})
 | 
|  |    125 | there is also the \emph{Archive of Formal Proofs}
 | 
|  |    126 | at  \url{http://afp.sourceforge.net}, a growing collection of Isabelle theories
 | 
|  |    127 | that everybody can contribute to.
 | 
|  |    128 | 
 | 
|  |    129 | \subsection{Quotation Marks}
 | 
|  |    130 | 
 | 
|  |    131 | The textual definition of a theory follows a fixed syntax with keywords like
 | 
|  |    132 | \isacommand{begin} and \isacommand{datatype}.  Embedded in this syntax are
 | 
|  |    133 | the types and formulae of HOL.  To distinguish the two levels, everything
 | 
|  |    134 | HOL-specific (terms and types) must be enclosed in quotation marks:
 | 
|  |    135 | \texttt{"}\dots\texttt{"}. To lessen this burden, quotation marks around a
 | 
|  |    136 | single identifier can be dropped.  When Isabelle prints a syntax error
 | 
|  |    137 | message, it refers to the HOL syntax as the \concept{inner syntax} and the
 | 
|  |    138 | enclosing theory language as the \concept{outer syntax}.
 | 
|  |    139 | \begin{warn}
 | 
|  |    140 | For reasons of readability, we almost never show the quotation marks in this
 | 
|  |    141 | book. Consult the accompanying theory files to see where they need to go.
 | 
|  |    142 | \end{warn}
 | 
|  |    143 | *}
 | 
|  |    144 | (*<*)
 | 
|  |    145 | end
 | 
|  |    146 | (*>*) |