(*****************************************************************
Isabelle Tutorial
-----------------
27th May 2009, Beijing
This file contains most of the material that will be covered in the
tutorial. The file can be "stepped through"; though it contains much
commented code (purple text).
*)
(*****************************************************************
Proof General
-------------
Proof General is the front end for Isabelle.
Proof General "paints" blue the part of the proof-script that has already
been processed by Isabelle. You can advance or retract the "frontier" of
this processed part using the "Next" and "Undo" buttons in the
menubar. "Goto" will process everything up to the current cursor position,
or retract if the cursor is inside the blue part. The key-combination
control-c control-return is a short-cut for the "Goto"-operation.
Proof General gives feedback inside the "Response" and "Goals" buffers.
The response buffer will contain warning messages (in yellow) and
error messages (in red). Warning messages can generally be ignored. Error
messages must be dealt with in order to advance the proof script. The goal
buffer shows which goals need to be proved. The sole idea of interactive
theorem proving is to get the message "No subgoals." ;o)
*)
(*****************************************************************
X-Symbols
---------
X-symbols provide a nice way to input non-ascii characters. For example:
\<forall>, \<exists>, \<Down>, \<sharp>, \<And>, \<Gamma>, \<times>, \<noteq>, \<in>, \<subseteq>, \<dots>
They need to be input via the combination \<name-of-x-symbol>
where name-of-x-symbol coincides with the usual latex name of
that symbol.
However, there are some handy short-cuts for frequently used X-symbols.
For example
[| \<dots> \<lbrakk>
|] \<dots> \<rbrakk>
==> \<dots> \<Longrightarrow>
=> \<dots> \<Rightarrow>
--> \<dots> \<longrightarrow>
/\ \<dots> \<and>
\/ \<dots> \<or>
|-> \<dots> \<mapsto>
=_ \<dots> \<equiv>
*)
(*****************************************************************
Theories
--------
Every Isabelle proof-script needs to have a name and must import
some pre-existing theory. This will normally be the theory Main.
*)
theory Lec1
imports "Main"
begin
text {*****************************************************************
Types
-----
Isabelle is based on types including some polymorphism. It also includes
some overloading, which means that sometimes explicit type annotations
need to be given.
- Base types include: nat, bool, string
- Type formers include: 'a list, ('a\<times>'b), 'c set
- Type variables are written like in ML with an apostrophe: 'a, 'b, \<dots>
Types known to Isabelle can be queried using:
*}
typ nat
typ bool
typ string (* the type for alpha-equated lambda-terms *)
typ "('a \<times> 'b)" (* pair type *)
typ "'c set" (* set type *)
typ "'a list" (* list type *)
typ "nat \<Rightarrow> bool" (* the type of functions from nat to bool *)
(* These give errors: *)
(*typ boolean *)
(*typ set*)
text {*****************************************************************
Terms
-----
Every term in Isabelle needs to be well-typed (however they can have
polymorphic type). Whether a term is accepted can be queried using:
*}
term c (* a variable of polymorphic type *)
term "1::nat" (* the constant 1 in natural numbers *)
term 1 (* the constant 1 with polymorphic type *)
term "{1, 2, 3::nat}" (* the set containing natural numbers 1, 2 and 3 *)
term "[1, 2, 3]" (* the list containing the polymorphic numbers 1, 2 and 3 *)
term "(True, ''c'')" (* a pair consisting of true and the string "c" *)
term "Suc 0" (* successor of 0, in other words 1 *)
text {*
Isabelle provides some useful colour feedback about what are constants (in black),
free variables (in blue) and bound variables (in green). *}
term "True" (* a constant that is defined in HOL *)
term "true" (* not recognised as a constant, therefore it is interpreted as a variable *)
term "\<forall>x. P x" (* x is bound, P is free *)
text {* Every formula in Isabelle needs to have type bool *}
term "True"
term "True \<and> False"
term "True \<or> B"
term "{1,2,3} = {3,2,1}"
term "\<forall>x. P x"
term "A \<longrightarrow> B"
text {*
When working with Isabelle, one is confronted with an object logic (that is HOL)
and Isabelle's meta-logic (called Pure). Sometimes one has to pay attention
to this fact. *}
term "A \<longrightarrow> B" (* versus *)
term "A \<Longrightarrow> B"
term "\<forall>x. P x" (* versus *)
term "\<And>x. P x"
term "A \<Longrightarrow> B \<Longrightarrow> C" (* is synonymous with *)
term "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> C"
text {*****************************************************************
Inductive Definitions: The even predicate
-----------------------------------------
Inductive definitions start with the keyword "inductive" and a predicate name.
One also needs to provide a type for the predicate. Clauses of the inductive
predicate are introduced by "where" and more than two clauses need to be
separated by "|". Optionally one can give a name to each clause and indicate
that it should be added to the hints database ("[intro]"). A typical clause
has some premises and a conclusion. This is written in Isabelle as:
"premise \<Longrightarrow> conclusion"
"\<lbrakk>premise1; premise2; \<dots> \<rbrakk> \<Longrightarrow> conclusion"
If no premise is present, then one just writes
"conclusion"
Below we define the even predicate for natural numbers.
*}
inductive
even :: "nat \<Rightarrow> bool"
where
eZ[intro]: "even 0"
| eSS[intro]: "even n \<Longrightarrow> even (Suc (Suc n))"
text {*****************************************************************
Theorems
--------
A central concept in Isabelle is that of theorems. Isabelle's theorem
database can be queried using
*}
thm eZ
thm eSS
thm conjI
thm conjunct1
text {*
Notice that theorems usually contain schematic variables (e.g. ?P, ?Q, \<dots>).
These schematic variables can be substituted with any term (of the right type
of course). It is sometimes useful to suppress the "?" from the schematic
variables. This can be achieved by using the attribute "[no_vars]". *}
thm eZ[no_vars]
thm eSS[no_vars]
thm conjI[no_vars]
thm conjunct1[no_vars]
text {*
When defining the predicate eval, Isabelle provides us automatically with
the following theorems that state how the even predicate can be introduced
and what constitutes an induction over the predicate even. *}
thm even.intros[no_vars]
thm even.induct[no_vars]
text {*****************************************************************
Lemma / Theorem / Corollary Statements
--------------------------------------
Whether to use lemma, theorem or corollary makes no semantic difference
in Isabelle. A lemma starts with "lemma" and consists of a statement
("shows \<dots>") and optionally a lemma name, some type-information for
variables ("fixes \<dots>") and some assumptions ("assumes \<dots>"). Lemmas
also need to have a proof, but ignore this 'detail' for the moment.
*}
lemma even_double:
shows "even (2 * n)"
by (induct n) (auto)
lemma even_add_n_m:
assumes a: "even n"
and b: "even m"
shows "even (n + m)"
using a b by (induct) (auto)
lemma neutral_element:
fixes x::"nat"
shows "x + 0 = x"
by simp
text {*****************************************************************
Isar Proofs
-----------
Isar is a language for writing down proofs that can be understood by humans
and by Isabelle. An Isar proof can be thought of as a sequence of 'stepping stones'
that start with the assumptions and lead to the goal to be established. Every such
stepping stone is introduced by "have" followed by the statement of the stepping
stone. An exception is the goal to be proved, which need to be introduced with "show".
Since proofs usually do not proceed in a linear fashion, a label can be given
to each stepping stone and then used later to refer to this stepping stone
("using").
Each stepping stone (or have-statement) needs to have a justification. The
simplest justification is "sorry" which admits any stepping stone, even false
ones (this is good during the development of proofs). Assumption can be
"justified" using "by fact". Derived facts can be justified using
- by simp (* simplification *)
- by auto (* proof search and simplification *)
- by blast (* only proof search *)
If facts or lemmas are needed in order to justify a have-statement, then
one can feed these facts into the proof by using "using label \<dots>" or
"using theorem-name \<dots>". More than one label at the time is allowed.
Induction proofs in Isar are set up by indicating over which predicate(s)
the induction proceeds ("using a b") followed by the command "proof (induct)".
In this way, Isabelle uses default settings for which induction should
be performed. These default settings can be overridden by giving more
information, like the variable over which a structural induction should
proceed, or a specific induction principle such as well-founded inductions.
After the induction is set up, the proof proceeds by cases. In Isar these
cases can be given in any order, but must be started with "case" and the
name of the case, and optionally some legible names for the variables
referred to inside the case.
The possible case-names can be found by looking inside the menu "Isabelle ->
Show me -> cases". In each "case", we need to establish a statement introduced
by "show". Once this has been done, the next case can be started using "next".
When all cases are completed, the proof can be finished using "qed".
This means a typical induction proof has the following pattern
proof (induct)
case \<dots>
\<dots>
show \<dots>
next
case \<dots>
\<dots>
show \<dots>
\<dots>
qed
The four lemmas are by induction on the predicate machines. All proofs establish
the same property, namely a transitivity rule for machines. The complete Isar
proofs are given for the first three proofs. The point of these three proofs is
that each proof increases the readability for humans.
*}
text {*****************************************************************
1.) EXERCISE
------------
Remove the sorries in the proof below and fill in the correct
justifications.
*}
lemma
shows "even (2 * n)"
proof (induct n)
case 0
show "even (2 * 0)" sorry
next
case (Suc n)
have ih: "even (2 * n)" by fact
have eq: "2 * (Suc n) = Suc (Suc (2 * n))" sorry
have h1: "even (Suc (Suc (2 * n)))" sorry
show "even (2 * (Suc n))" sorry
qed
text {*****************************************************************
2.) EXERCISE
------------
Remove the sorries in the proof below and fill in the correct
justifications.
*}
lemma
shows "even (n + n)"
proof (induct n)
case 0
show "even (0 + 0)" sorry
next
case (Suc n)
have ih: "even (n + n)" by fact
have eq: "(Suc n) + (Suc n) = Suc (Suc (n + n))" sorry
have a: "even (Suc (Suc (n + n)))" sorry
show "even ((Suc n) + (Suc n))" sorry
qed
text {*
Just like gotos in the Basic programming language, labels can reduce
the readability of proofs. Therefore one can use in Isar the notation
"then have" in order to feed a have-statement to the proof of
the next have-statement. In the proof below this has been used
in order to get rid of the labels a.
*}
lemma even_twice:
shows "even (n + n)"
proof (induct n)
case 0
show "even (0 + 0)" by auto
next
case (Suc n)
have ih: "even (n + n)" by fact
have eq: "(Suc n) + (Suc n) = Suc (Suc (n + n))" by simp
have "even (Suc (Suc (n + n)))" using ih by auto
then show "even ((Suc n) + (Suc n))" using eq by simp
qed
text {*
The label eq cannot be got rid of in this way, because both
facts are needed to prove the show-statement. We can still avoid the
labels by feeding a sequence of facts into a proof using the chaining
mechanism:
have "statement1" \<dots>
moreover
have "statement2" \<dots>
\<dots>
moreover
have "statementn" \<dots>
ultimately have "statement" \<dots>
In this chain, all "statementi" can be used in the proof of the final
"statement". With this we can simplify our proof further to:
*}
lemma
shows "even (n + n)"
proof (induct n)
case 0
show "even (0 + 0)" by auto
next
case (Suc n)
have ih: "even (n + n)" by fact
have "(Suc n) + (Suc n) = Suc (Suc (n + n))" by simp
moreover
have "even (Suc (Suc (n + n)))" using ih by auto
ultimately show "even ((Suc n) + (Suc n))" by simp
qed
text {*
While automatic proof procedures in Isabelle are not able to prove statements
like "P = NP" assuming usual definitions for P and NP, they can automatically
discharge the lemmas we just proved. For this we only have to set up the induction
and auto will take care of the rest. This means we can write:
*}
lemma
shows "even (2 * n)"
by (induct n) (auto)
lemma
shows "even (n + n)"
by (induct n) (auto)
text {*****************************************************************
Rule Inductions
---------------
Inductive predicates are defined in terms of rules of the form
prem1 \<dots> premn
--------------
concl
In Isabelle we write such rules as
\<lbrakk>prem1; \<dots>; premn\<rbrakk> \<Longrightarrow> concl
You can do induction over such rules according to the following
scheme:
1.) Assume the property for the premises. Assume the side-conditions.
2.) Show the property for the conclusion.
In Isabelle the corresponding theorem is called, for example
*}
thm even.induct
text {*****************************************************************
3.) EXERCISE
------------
Remove the sorries in the proof below and fill in the correct
justifications.
*}
lemma even_add:
assumes a: "even n"
and b: "even m"
shows "even (n + m)"
using a b
proof (induct)
case eZ
have as: "even m" by fact
show "even (0 + m)" sorry
next
case (eSS n)
have ih: "even m \<Longrightarrow> even (n + m)" by fact
have as: "even m" by fact
show "even (Suc (Suc n) + m)" sorry
qed
text {*
Whenever a lemma is of the form
lemma
assumes "pred"
and \<dots>
shows "something"
then a rule induction is generally appropriate (but
not always).
In case of even_add it is definitely *not* appropriate.
Compare for example what you have to prove when you attempt
an induction over natural numbers.
*}
lemma even_add_does_not_work:
assumes a: "even n"
and b: "even m"
shows "even (n + m)"
using a b
proof (induct n rule: nat_induct)
case 0
have "even m" by fact
then show "even (0 + m)" by simp
next
case (Suc n)
have ih: "\<lbrakk>even n; even m\<rbrakk> \<Longrightarrow> even (n + m)" by fact
have as1: "even (Suc n)" by fact
have as2: "even m" by fact
show "even ((Suc n) + m)" oops
text {*****************************************************************
4.) EXERCISE (Last fact about even)
-----------------------------------
Remove the sorries in the proof below and fill in the correct
justifications. The following two lemmas should be useful
*}
thm even_twice[no_vars]
thm even_add[no_vars]
text {*
Remember you can include lemmas in the "using" statement,
just like other facts. For example:
have "something" using even_twice by simp
*}
lemma even_mul:
assumes a: "even n"
shows "even (n * m)"
using a
proof (induct)
case eZ
show "even (0 * m)" by auto
next
case (eSS n)
have as: "even n" by fact
have ih: "even (n * m)" by fact
show "even (Suc (Suc n) * m)" sorry
qed
text {*****************************************************************
Definitions
-----------
Often it is useful to define new concepts in terms of existsing
concepts. In Isabelle you introduce definitions with the key-
word "definition". For example
*}
definition
divide :: "nat \<Rightarrow> nat \<Rightarrow> bool" ("_ DVD _" [100,100] 100)
where
"m DVD n = (\<exists>k. n = m * k)"
text {*
The annotation ("_ DVD _" [100,100] 100) introduces some fancier
syntax. In this case we define this predicate as infix. The underscores
correspond to the to arguments. The numbers stand for precedences.
Once this definition is stated, it can be accessed as a normal
theorem. For example
*}
thm divide_def
text {*
Below we prove the really last fact about even. Note
how we deal with existential quantified formulas, where
we have to use
obtain \<dots> where \<dots>
in order to obtain a witness with which we can reason further.
*}
lemma even_divide:
assumes a: "even n"
shows "2 DVD n"
using a
proof (induct)
case eZ
have "0 = 2 * (0::nat)" by simp
then show "2 DVD 0" by (auto simp add: divide_def)
next
case (eSS n)
have "2 DVD n" by fact
then have "\<exists>k. n = 2 * k" by (simp add: divide_def)
then obtain k where eq: "n = 2 * k" by (auto)
have "Suc (Suc n) = 2 * (Suc k)" using eq by simp
then have "\<exists>k. Suc (Suc n) = 2 * k" by blast
then show "2 DVD (Suc (Suc n))" by (simp add: divide_def)
qed
text {*****************************************************************
Function Definitions
--------------------
Many functions over datatypes can be defined by recursion on the
structure. For this purpose, Isabelle provides "fun". To use it one needs
to give a name for the function, its type, optionally some pretty-syntax
and then some equations defining the function. Like in "inductive",
"fun" expects that more than one such equation is separated by "|".
Below we define function iteration using function composition
defined as "o".
*}
fun
iter :: "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> ('a \<Rightarrow> 'a)" ("_ !! _")
where
"f !! 0 = (\<lambda>x. x)"
| "f !! (Suc n) = (f !! n) o f"
text {*
5.) EXERCISE
Prove the following property by induction over n.
*}
lemma
shows "f !! (m + n) = (f !! m) o (f !! n)"
sorry
text {*
A textbook proof of this lemma usually goes:
Case 0: Trivial
Case (Suc n): We have to prove that
"f !! (m + (Suc n)) = f !! m o (f !! (Suc n))"
holds. The induction hypothesis is
"f !! (m + n) = (f !! m) o (f !! n)"
The proof is as follows
"f !! (m + (Suc n)) = f !! (Suc (m + n))"
= f !! (m + n) o f
= (f !! m) o (f !! n) o f (by ih)
= (f !! m) o ((f !! n) o f) (by o_assoc)
= (f !! m) o (f !! (Suc n))
Done.
*}
lemma
shows "f !! (m + n) = (f !! m) o (f !! n)"
proof (induct n)
case 0
show "f !! (m + 0) = (f !! m) o (f !! 0)" sorry
next
case (Suc n)
have ih: "f !! (m + n) = (f !! m) o (f !! n)" by fact
show "f !! (m + (Suc n)) = f !! m o (f !! (Suc n))" sorry
qed
text {*
The following proof involves several steps of equational reasoning.
Isar provides some convenient means to express such equational
reasoning in a much cleaner fashion using the "also have"
construction. *}
lemma
shows "f !! (m + n) = (f !! m) o (f !! n)"
proof (induct n)
case 0
show "f !! (m + 0) = (f !! m) o (f !! 0)" by (simp add: comp_def)
next
case (Suc n)
have ih: "f !! (m + n) = (f !! m) o (f !! n)" by fact
have eq1: "f !! (m + (Suc n)) = f !! (Suc (m + n))" by simp
have eq2: "f !! (Suc (m + n)) = f !! (m + n) o f" by simp
have eq3: "f !! (m + n) o f = (f !! m) o (f !! n) o f" using ih by simp
have eq4: "(f !! m) o (f !! n) o f = (f !! m) o ((f !! n) o f)" by (simp add: o_assoc)
have eq5: "(f !! m) o ((f !! n) o f) = (f !! m) o (f !! (Suc n))" by simp
show "f !! (m + (Suc n)) = f !! m o (f !! (Suc n))"
using eq1 eq2 eq3 eq4 eq5 by (simp only:)
qed
text {*
The equations can be stringed together with "also have" and "\<dots>",
as shown below. *}
lemma
shows "f !! (m + n) = (f !! m) o (f !! n)"
proof (induct n)
case 0
show "f !! (m + 0) = (f !! m) o (f !! 0)" by (simp add: comp_def)
next
case (Suc n)
have ih: "f !! (m + n) = (f !! m) o (f !! n)" by fact
have "f !! (m + (Suc n)) = f !! (Suc (m + n))" by simp
also have "\<dots> = f !! (m + n) o f" by simp
also have "\<dots> = (f !! m) o (f !! n) o f" using ih by simp
also have "\<dots> = (f !! m) o ((f !! n) o f)" by (simp add: o_assoc)
also have "\<dots> = (f !! m) o (f !! (Suc n))" by simp
finally show "f !! (m + (Suc n)) = f !! m o (f !! (Suc n))" by simp
qed
text {*
This type of equational reasoning also extends to relations.
For this consider the following power function and the auxiliary
fact about less-equal and multiplication. *}
fun
pow :: "nat \<Rightarrow> nat \<Rightarrow> nat" ("_ \<up> _")
where
"m \<up> 0 = 1"
| "m \<up> (Suc n) = m * (m \<up> n)"
lemma aux:
fixes a b c::"nat"
assumes a: "a \<le> b"
shows " (c * a) \<le> (c * b)"
using a by (auto)
text {*
With this you can easily implement the following proof
which is straight taken out of Velleman's "How To Prove It".
*}
lemma
shows "1 + n * x \<le> (1 + x) \<up> n"
proof (induct n)
case 0
show "1 + 0 * x \<le> (1 + x) \<up> 0" by simp
next
case (Suc n)
have ih: "1 + n * x \<le> (1 + x) \<up> n" by fact
have "1 + (Suc n) * x \<le> 1 + x + (n * x) + (n * x * x)" by (simp)
also have "\<dots> = (1 + x) * (1 + n * x)" by simp
also have "\<dots> \<le> (1 + x) * ((1 + x) \<up> n)" using ih aux by blast
also have "\<dots> = (1 + x) \<up> (Suc n)" by simp
finally show "1 + (Suc n) * x \<le> (1 + x) \<up> (Suc n)" by simp
qed
text {*
What Velleman actually wants to prove is a proper
inequality. For this we can nest proofs.
*}
lemma
shows "n * x < (1 + x) \<up> n"
proof -
have "1 + n * x \<le> (1 + x) \<up> n"
proof (induct n)
case 0
show "1 + 0 * x \<le> (1 + x) \<up> 0" by simp
next
case (Suc n)
have ih: "1 + n * x \<le> (1 + x) \<up> n" by fact
have "1 + (Suc n) * x \<le> 1 + x + (n * x) + (n * x * x)" by (simp)
also have "\<dots> = (1 + x) * (1 + n * x)" by simp
also have "\<dots> \<le> (1 + x) * ((1 + x) \<up> n)" using ih aux by blast
also have "\<dots> = (1 + x) \<up> (Suc n)" by simp
finally show "1 + (Suc n) * x \<le> (1 + x) \<up> (Suc n)" by simp
qed
then show "n * x < (1 + x) \<up> n" by simp
qed
end