SlideShare una empresa de Scribd logo
1 de 28
Descargar para leer sin conexión
A Machine-Assisted Proof of
Gödel's Incompleteness Theorems
Lawrence C. Paulson, Computer Laboratory, University of Cambridge
The most misunderstood
theorems in mathematics
✤

Gödel’s theorems have highly
technical, syntactic proofs.
1. Every “strong enough”
formal system is incomplete,
in that at least one formula
can neither be proved nor
disproved.
2. And if such a formal system
admits a proof of its own
consistency, then it is actually
inconsistent.

✤

For the first time, both of
Gödel’s proofs have been
mechanised, following a paper
by S!wierczkowski (2003)

✤

The machine proof, in the
structured Isar language, is
complete, almost readable, and
can be perused interactively.
Hereditarily finite set theory
✤

A hereditarily finite set is a finite set of HF sets.

✤

Many mathematical constructions, including natural numbers and
sequences, can be defined as in standard set theory.

✤

HF set theory is equivalent to Peano Arithmetic via the mapping

f(x) =

X

f(y)

{2

| y 2 x}
Benefits of Using HF Set Theory
✤

Can use standard definitions of
pairing and sequences.

✤

The first incompleteness
theorem requires an HF
development of the natural
numbers, induction, etc., but
not addition.

✤

The second incompleteness
theorem requires operations on
sequences and addition, but not
multiplication.

✤

No need for least common
multiples, prime numbers or
the Chinese remainder theorem.
The Axioms of HF Set Theory

✤

0 denotes the empty set

✤

x ◁ y denotes the set x extended with the element y.

✤

There are no other function symbols.

✤

Union, intersection, etc can be shown to exist by induction.
Stages of the Proofs
✤

The syntax of a first-order
theory is formalised: terms,
formulas, substitution...

✤

A system of coding to formalise
the calculus within itself. The
code of # is a term, written ⌜#⌝.

✤

A deductive calculus for sequents
of the form " ⊦ # (typically for
Peano arithmetic, but here HF)

✤

Syntactic predicates to
recognise codes of terms,
substitution, axioms, etc.

✤

Meta-theory to relate truth and
provability. E.g. “all true $
formulas are theorems”. $
formulas are built using % & ∃
and bounded ∀.

✤

Finally the predicate Pf, such
that ⊦ # ⟺ ⊦ Pf ⌜#⌝.
First Incompleteness Theorem
✤

To prove Gödel’s first incompleteness theorem, construct ' that
expresses that ' is not provable.

✤

It follows (provided the calculus is consistent) that neither ' nor its
negation can be proved.

✤

Need to show that substitution behaves like a function.
✤

Requires a detailed proof in the calculus,

✤

… alternatively, other detailed calculations.
Second Incompleteness Theorem

✤

This crucial lemma for Gödel’s second incompleteness theorem is
proved by induction over the construction of # as a $ formula.

✤

It requires generalising the statement above to allow the formula # to
contain free variables.
✤

complex technicalities

✤

lengthy deductions in the calculus
Proving Theorems in the Calculus
✤

Gödel knew that formal proofs
were difficult.They could be
eliminated, but at what cost?

✤

By coding all predicates as
executable functions, and
proving a meta-theorem, Gödel
reduced provability to truth.

✤

But then only bounded
quantifiers can be used, with
tricky arithmetical proofs that
the bounds are adequate.

✤

With $ formulas, provability is
reduced to truth for most
desired properties, with no
tricky proofs about bounds.

✤

Instead, some straightforward
inductions need to be
formalised in the calculus.

✤

The second theorem requires
working in the calculus
anyway.
Isabelle/HOL and Nominal
✤

a proof assistant for higherorder logic

✤

Free names are significant, but
not once they are bound.

✤

much automation to hide the
underlying proof calculus

✤

✤

support for recursive functions
and inductive sets

Syntax involving variable
binding can be defined using
recursion, provided variables
are used “sensibly”.

✤

During proof by induction,
bound variable names can be
guaranteed not to clash with
specified other terms.

✤

the nominal package, for working
with named variables
De Bruijn Indexes
✤

This approach to variable binding replaces names by numbers.

✤

0 denotes the innermost bound variable, 1 for the next, etc.

✤

This approach destroys readability, but substitution and abstraction
are very easy to define.

✤

During coding, formulas are translated into the de Bruijn format.

✤

And so there is no need to formalise the nominal theory within HF.
1

The Isabelle/HOL proof development: fundamentals.

Defining Terms and Formulas
z = 0 $ 8x [x 62 z]
z = x C y $ 8u [u 2 z $ u 2 x _ u = y]
(0) ^ 8xy [ (x) ^ (y) ! (x C y)] ! 8x [ (x)]

(HF1)
(HF2)
(HF3)

nominal datatype tm = Zero | Var name | Eats tm tm
nominal datatype fm =
Mem tm tm
( infixr "IN"
| Eq tm tm
( infixr "EQ"
| Disj fm fm
( infixr "OR"
| Neg fm
| Ex x::name f::fm binds x

150)
150)
130)
in f

The HF calculus includes an existential quantifier, denoted Ex, which involves
variable binding via the nominal framework. The infixr declarations provide
an alternative syntax for the membership relation, the equality relation, and
disjunction. A formula can also be a negation. The other logical connectives
Variable binding
are introduced later as abbreviations.
formalised to formalise, but here it is straightSubstitution is often problematicalusing nominal
forward. Substitution of a term x for a variable i is defined as follows:
nominal primrec subst :: "name ) tm ) tm ) tm"
where
"subst i x Zero
= Zero"
The HF calculus includes an existential quantifier, denoted Ex, which involves
variable binding via the nominal framework. The infixr declarations provide
an alternative syntax for the membership relation, the equality relation, and
disjunction. A formula can also be a negation. The other logical connectives
are introduced later as abbreviations.
Substitution is often problematical to formalise, but here it is straightforward. Substitution of a term x for a variable i is defined as follows:

Defining Substitution

nominal primrec subst :: "name ) tm ) tm ) tm"
where
"subst i x Zero
= Zero"
For substitution within a formula, we normally expect issues concerning
| "subst i x (Var k)
= (if i=k then x else Var k)"
the "subst iofx a(Eats t variable. Note that thet) (subst substituting the
bound u) = Eats (subst i x result of i x u)"
| capture
term x for the variable i in the formula A is written A(i::=x).

1
nominal primrec subst fm :: "fm ) name ) tm ) fm"
where
Mem: "(Mem t u)(i::=x) = Mem (subst i x t) (subst i x u)"
| Eq:
"(Eq t u)(i::=x)
= Eq (subst i x t) (subst i x u)"
| Disj: "(Disj A B)(i::=x) = Disj (A(i::=x)) (B(i::=x))"
| Neg: "(Neg A)(i::=x)
= Neg (A(i::=x))"
| Ex:
"atom j ] (i, x) =) (Ex j A)(i::=x) = Ex j (A(i::=x))"
Substitution is again straightforward in the first four cases (membership,
Properties of
equality, disjunction, negation). be the existential case, the precondition
In
The variable j must
atom j ] (i, x) (pronounced “j is fresh for i and xsubstitution have
”) essentially says that
fresh di↵erent names
i and j must be for i and x with j not free in x. We do not need to
simple proofs.
supply a mechanism for renaming the bound variable, as that is part of the
nominal framework, which in most cases will choose a su ciently fresh bound
supply a mechanism for renaming the bound variable, as that is part of the
nominal framework, which in most cases will choose a su ciently fresh bound
variable at the outset. The usual properties of substitution (commutativity,
for example) have simple proofs by induction on formulas. In contrast, ?)
needed to combine three substitution lemmas in a simultaneous proof by
induction, a delicate argument involving 1900 lines of Coq.
The HF proof system is an inductively defined predicate, where H ` A
means that the formula A is provable from the set of formulas H .

Defining the HF Calculus

inductive hfthm :: "fm set ) fm ) bool" ( infixl " `" 55)
where
Hyp:
"A 2 H =) H ` A"
| Extra: "H ` extra axiom"
| Bool:
"A 2 boolean axioms =) H ` A"
| Eq:
"A 2 equality axioms =) H ` A"
| Spec:
"A 2 special axioms =) H ` A"
| HF:
"A 2 HF axioms =) H ` A"
| Ind:
"A 2 induction axioms =) H ` A"
| MP:
"H ` A IMP B =) H’ ` A =) H [ H’ ` B"
| Exists: "H ` A IMP B =)
atom i ] B =) 8 C 2H. atom i ] C =) H ` (Ex i A) IMP B"

Note that the existential rule is subject to the condition that the bound
variable, i, is fresh with respect to B and the formulas in H . The definitions
The variable i must be
of boolean axioms, etc., are taken from ?). He formalised a simpler inference
system, with theorems of the H
fresh for B and form ` A, but introducing H allows a proof of
the deduction theorem and the derivation of a sort of sequent calculus. This
is essential if we are to conduct proofs in this formal calculus.
Early Steps in the HF Calculus
✤

the deduction theorem (yielding a sequent calculus)

✤

derived rules to support explicit formal proofs
✤

✤

✤

for defined connectives, including ∧ ➝ ∀
for equality, set induction, …

definitions and proofs for subsets, extensionality, foundation and natural
number induction
states that the quantified variable (z ) must be fresh for the terms t and u.
MemI: and in contrast to some treatments, the bound variable is a
In other words,"ss fm (Var i IN Var j)"
| DisjI: "ss fm A =) ss than being fm (A OR B)"
parameter of the definition ratherfm B =) ssfixed; however, the choice of z
| a↵ect "ss fm A =) ss fm B =) ss fm (A AND to
cannot ConjI:the denotation of the right-hand side, thanksB)" quotienting.
| ExI:
"ss fm A =) ss fm (Ex i A)"
| All2I: "ss "fm =) atom j ] (i,A) =) ss fm (All2 i (Var j) A)"
inductive ss fm ::fm A ) bool" where

Σ Formulas

MemI: "ss fm (Var i IN Var j)"
One advantage A =) ss fm B =) ss fm (A OR B)"
| DisjI: "ss fmof formal proof is that these conditions are immediately ev| ConjI: "sstheyAmay not fm B =) from an informal presentation. ?) does
ident, when fm =) ss be clear ss fm (A AND B)"
| ExI:
"ss fm A condition (Ex i A)"
not impose the last =) ss fm (on the bound of a universal quantifier), but it
statesStrict fm A =) atom j ] containto reach the second j) A)"
terms t the
| All2I: "ss $ quantified variable(i,A) =) bevariablesiand areand u.basis
fm (All2
greatlythat the formulas only(z )needed ssfresh for the (Var incompletesimplifies the main induction must

for
In other words, and are only interested treatments, the boundincompleteness
ness theorem. (If we in contrast to somein second incompleteness theorem.
formalising the first variable is a
the the definition is that the
parameter of of formal proofrather than being fixed; are immediately evOne advantagemain induction of these conditionshowever, the choice of z
theorem, we can use a more generous notion of ⌃ formula, allowing atomic
ident, when they may notstill derive the general caseto quotienting.
be clear from an informal thanks of ?) does
cannot a↵ect the denotation of the right-hand side,presentation. $ formulas.
formulas andWe can
their negations over arbitrary terms.) Formally, a ⌃ formula is
not impose the last condition (on the bound of a universal quantifier), but it
defined to be any formula) bool" where
inductive ss the main that can be proved equivalent (in the HF calculus)
greatly simplifiesfm :: "fminduction needed to reach the second incompleteto a strict ⌃"ss fm (Var interested in
MemI: formula:
ness theorem. (If we are onlyi IN Var j)"formalising the first incompleteness
| DisjI:
OR B)"
theorem, we fm "ss fma Amore generous =) ss of ⌃✓formula, allowing atomic B)"
can use! ( =) ss fm BB notion fm (A supp A & {} ` A IFF
"Sigma
A
9 B. ss fm
& supp B
| ConjI:their negations over arbitrary ss fm (AFormally, a ⌃ formula is
"ss fm A =) ss fm B =) terms.) AND B)"
formulas and
| to be "ss fm
The condition suppAB that can A(Ex i A)"
✓ ss fm proved equivalent that every variable
defined ExI: any formula=) supp be essentially means (in the HF calculus) free
to in strict ⌃ also be free =)A . After a certain =) ss fm (All2 iit(Var j) A)"to
a | All2I: "ss fm A in atom j ] (i,A) amount of e↵ort, is possible
B must formula:
derive the expected properties of ⌃ formulas and ultimately to reach a key
One advantage of formal proof & supp B ✓ supp A & are immediately
"Sigma fm A ! ( 9 B. ss fm B is that these conditions {} ` A IFF B)" evresult based they may not be clear from an informal presentation. ?) does
on this concept:
ident, when
The condition the last conditionA(on the bound of a universal quantifier),free it
not impose supp B ✓ supp essentially means that every variable
theorem " [[Sigma fm A; ground fm A; eval fm e0 A ]] =) {} ` A" but
in B must also be free inmain induction needed to reach the second incompletegreatly simplifies the A . After a certain amount of e↵ort, it is possible to
derive the expected we are only of ⌃ formulas and ultimately to incompleteness `
If A is a true ⌃ sentence, then `A. This result reduces the task of a key
ness theorem. (If properties interested in formalising the first reach proving
result basedformal calculus to proving that A holds formula, allowing atomic in
A in the we can use a more generous notion of ⌃ (written eval fm e0 A )
theorem, on this concept:

True $ formulas are theorems!

formulas and their negations over fm A;
Isabelle/HOL’s imp thm: " [[Sigmaarbitrary terms.) A; eval a e0 A ]] =)
theorem Sigma fm native higher-order logic.ground fm Formally, fm ⌃ formula is
Coding Terms and Formulas
✤

must first translate from nominal to de Bruijn format
✤

✤

✤

the actual coding is a simple recursive map:
⌜0⌝=0 , ⌜xk⌝=k, ⌜x ◁ y⌝= 〈⌜◁⌝, ⌜x⌝, ⌜y⌝〉, …

also define (in HF) predicates to recognise codes
✤

abstraction over a variable — needed to define Form(x), the
predicate for formulas

✤

substitution for a variable
details. Many other authors prefer to simplify matters via repeated appeals
where "MakeForm y u w ⌘
to Church’s thesis. Even the detailed presentations mentioned above omit any
y = q Disj u w _ y = q Neg u _
demonstration that the definitions are correct. The proof formalisation condi( 9 v u’. AbstForm v 0 u u’ ^ y = q Ex u’)"
tion for the provability predicate (written PfP below) is typically stated with a
minimum of code of a formula constructed from existing formulas u and v
Thus y is thejustification:
provided y codes the disjunction u _ v, the negation ¬u or the existential formula
theorem proved iff proved Pf: "{} ` ↵ ! {} ` PfP p↵q"
9(u’), where u’ has been obtained by abstracting u over some variable, v. The
predicate interplay of these various points can be seen below:
The AbstForm performs de Bruijn abstraction over a formula; its definition
is complicated, and omitted here. Note that the codes of quantified formulas do
definition the names
not mention MakeForm ::of"hf ) hf ) hf ) bool"
bound variables.
where "MakeForm is given ⌘ a higher-order logic formula, and therefore at the
This predicate y u w by
y = q Disj u w _ y = q Neg u _
level of the meta-theory. Working at this level eliminates the need to construct
( 9 v u’. AbstForm v 0 u u’ ^ y = q Ex u’)"
HF proofs, and most of the correctness properties we need can be proved in
this manner. However, in formula constructed from existing formulas u and v
Thus y is the code of a % w, or y = ¬ u, ordiagonalisation argument and
y = u order to perform the y = (∃v) u
exhibit theyundecidable formula, we need a negationof u or the existential formula
provided codes the disjunction u _ v, the version ¬ every coding predicate as
with an explicit abstraction over on both levels:
an HF formula. Therefore, obtained by abstracting definedon u variable, v. The
9(u’), where u’ has been each predicate must be ustep some
predicate AbstForm performs de Bruijn abstraction over a formula; its definition
nominal primrec MakeFormP :: "tm ) tm ) tm ) fm"
is complicated, and omitted here. Note that the codes of quantified formulas do
where " [[atom v ] (y,u,w,au); atom au ] (y,u,w) ]] =)
not mentionythe wnames of bound variables.
MakeFormP
u
=
This predicate is given by higher-order logic formula, and therefore at the
y EQ Q Disj u w OR y EQ QaNeg u OR
level of the meta-theory. Workingv) Zero level eliminates theEQ Q Ex (Var au)))"
Ex v (Ex au (AbstFormP (Var at this u (Var au) AND y need to construct
HF proofs, and most of the correctness properties we need can be proved in
this manner. However, in order to perform the diagonalisation argument and
The “official” version as an HF formula, every boolean
exhibit the undecidable formula, we need a version of not acoding predicate as
an HF formula. Therefore, each predicate must be defined on both levels:
3

Example: Making a Formula

nominal primrec MakeFormP :: "tm ) tm ) tm ) fm"
Those Coding Predicates
SeqTerm

Term

SeqStTerm
AbstAtomic

VarNonOccTerm

AbstTerm
SeqAbstForm

SubstAtomic
Atomic

SeqConst

SubstTerm
AbstForm

SeqSubstForm
MakeForm

VarNonOccAtomic

Const

SubstForm
SeqForm

SeqVarNonOccForm

Form
VarNonOccForm
… And Proof Predicates
Sent

Equality_ax
Axiom
Prf

HF_ax

ModPon

Special_ax
Exists

Induction_ax
Subst

(a sequence of proof steps, and finally...)

Pf
quantified variables. Here there are only two, but to define AbstForm requires
As bound variables. the definition of Subset, constraints are and messy, and
12 we saw above in The necessary declarations are lengthy required on all
quantified variables. on thethere are only two, but torun very slowly), requires
Here
put a heavy burdenabove innominal package (proofs , define AbstForm but the on all
As we saw The necessary declarations are constraints messy, and
the definition of
12 bound variables. to rename explicit bound Subset lengthy and are required
alternative of having
variables
unattractive.
quantified variables.nominal package only two, but alsoslowly), but the requires
Here there are (proofs run is to define AbstForm
put weheavy above inon the
burden the definition of Subset, constraints are required on all
As a sawbound variables. The necessary declarations very lengthy and messy, and
12 of having
alternative variables. to rename explicit bound variables isare AbstForm requires
also
quantified a heavy burden on the nominal package (proofs run very slowly), but the
Here there are only two, but to define unattractive.
put
3 bound variables. having to rename explicit bound lengthy andalso unattractive.
proof development: first
12 The Isabelle/HOL declarations are variables is messy, and
alternative of The necessary
put incompleteness theorem. development: first
heavy Isabelle/HOL proof
3 aThe burden on the nominal package (proofs run very slowly), but the
alternative of having to rename explicit bound variables is also unattractive.
incompleteness K(p q) = p (pproof development: first
3 The Isabelle/HOL q)q
` theorem.

Steps to the First Theorem

incompleteness theorem.
first
We lemmaaprove KRP: "{}such that i qK(pq q) = stated, but it is neither a senneedincompleteness pVar ` pA easily p (p q) q"
function K ` KRP theorem.
The property of being single-valued is pA(i::= pA q)q
` pVar
p pA(i::=
3 The KRP: "{} ` KRP K(p i q=pA q (p development:
lemma prove Isabelle/HOLq)proof q)q pA q) q"

✤

✤

✤

✤

tence nor a ⌃ formula. Proving`thisq) = p (pq q)qq pA(i::=600 lines of explicit
lemmano function symbols. Instead, defineit q) q"
prove KRP:single-valuedpVar requires about pA is neither a KRP:
"{}K(p result i pA
KRP is easily stated, but a relation, sen… but The propertyin the HF calculus, verifying that substitution over terms or
we have of being `
reasoning steps
tence nor a ⌃ formula. Proving this result requires about 600 lines of explicit
formulasprove KRP: "{} result. single-valued is easily stated,
lemma yields a unique
The in the HF calculus, verifying that pA q) q"
being
reasoning steps property`ofKRP pVar i q pA q pA(i::=substitution but it terms or a senover is neither
tence nor a
Proving
requires
formulasKRP unique: ⌃ formula. y, KRP v this resulty’ proof about 600 lines of explicit
yields a
lemma property unique result.
"{KRP single-valued isx easily`
v x
y’} HF EQ but it is neither a senProving its functional behaviour takes 600 stated, y" steps.
The reasoning of being the HF calculus, verifying that substitution over terms or
steps in
tence nor a unique: "{KRP now reached by y’} ` y’ about 600 lines of explicit
lemma formulas lemmaaProving result.v xthe standard argument. The obtains
v x y, KRP
The KRP ⌃ formula. isunique this result requires EQ y"
diagonal yields
reasoning steps in form of calculus, quantification, and is equivalent terms or
syntax represents a the HF existential verifying that substitution over to 9 ....
Finally, the yields alemma isresult. v x by the standard argument.y" obtains
The diagonalKRP unique: "{KRP
lemma unique now
formulas diagonal lemma:reached y, KRP v x y’} ` y’ EQ The
syntax represents a form of existential quantification, and is equivalent to 9 ....
lemma diagonal:
lemma KRPThe diagonal lemma y, now reached by the standard argument. The obtains
x is KRP v x y’} "supp
obtains unique: "{} ` v IFF ↵(i::= p q)" ` y’ EQ y" supp ↵ - {atom i}"
where "{KRP
=
lemma syntax represents a form of existential quantification, and is equivalent to 9 ....
diagonal:
obtains part lemma is now reached by q)" "supp
supp
- i}, states
TheThe diagonal of the conclusion, namelythe standard argument.↵The obtains
second where "{} ` IFF ↵(i::= p supp = supp= ↵ - {atom {atom i}"
lemma diagonal:
syntax represents a form the formula are those of ↵ and the exception 9 ....
that second part of theof of existential quantification, supp is↵equivalenti},of i ; it
the free variables conclusion, namely supp = with - {atom to states
The
obtains
where "{} `
IFF ↵(i::= p q)" "supp
= supp ↵ - {atom i}"
is necessary in order to show that the undecidable formula is actually a sentence.
theorem Goedel I:
assumes Con: " ¬ {} ` Fls"
obtains
where "{} `
IFF Neg (PfP p q)"
" ¬ {} ` " " ¬ {} ` Neg "
"eval fm e " "ground fm "
proof obtain
where
"{} `
IFF Neg ((PfP (Var i))(i::= p q))"
and [simp]: "supp
= supp (Neg (PfP (Var i))) - {atom i}"
by (metis SyntaxN.Neg diagonal)
hence diag: "{} `
IFF Neg (PfP p q)"
by simp
hence np: " ¬ {} ` "
by (metis Con Iff MP same Neg D proved iff proved Pf)
hence npn: " ¬ {} ` Neg " using diag
by (metis Iff MP same NegNeg D Neg cong proved iff proved Pf)
moreover have "eval fm e " using hfthm sound [ where e=e, OF diag]
by simp (metis Pf quot imp is proved np)
moreover have "ground fm "
by (auto simp: ground fm aux def)
ultimately show ?thesis
by (metis diag np npn that)
qed
Steps to the Second Theorem
✤

Coding must be generalised to allow variables in codes.
✤

✤

✤

✤

⌜x ◁ y⌝ = 〈⌜◁⌝, ⌜x⌝, ⌜y⌝〉
⎣x ◁ y⎦V = 〈⌜◁⌝, x, y〉

codes of variables
are integers

Variables must be renamed, with the intent of creating “pseudoterms” like 〈⌜◁⌝, Q x, Q y〉.
Q is a magic function: Q x = ⌜t⌝ where t is some canonical term
denoting the set x.
Complications

✤

Q must be a relation.
✤

Function symbols cannot be added…

✤

Sets do not have an easily defined canonical ordering.

✤

QR(0,0)

✤

QR(x,x’), QR(y,y’) ⟹ QR(x ◁ y, 〈⌜◁⌝, x’, y’〉)
?) eliminates the ambiguity implicit in (5) above by appealing to a total
ordering, <, on the HF universe. The di culty is how to define this ordering
´
within the HF calculus. Swierczkowski develops the theory, including a definition by recursion on the rank of a set, but it does not look easy to formalise
in HF. P
Another approach is to define the function f : HA ! N such that
f (x) =
{2f (y) | y 2 x}. Then we can define x < y () f (x) < f (y). Again,
the e↵ort to formalise this theory in HF may be simpler than that needed to
formalise the Chinese remainder theorem, but it is still considerable.

One of the Final Lemmas
QR(x, x0 ), QR(y, y 0 ) ` x 2 y ! Pf bx0 2 y 0 c{x0 ,y0 }

QR(x, x0 ), QR(y, y 0 ) ` x ✓ y ! Pf bx0 ✓ y 0 c{x0 ,y0 }
QR(x, x0 ), QR(y, y 0 ) ` x = y ! Pf bx0 = y 0 c{x0 ,y0 }

✤

✤

2 ; () ?
The first two require zsimultaneous induction, yielding the third.
z 2 x C y () z 2 x _ z = y

; symbols %
Similar proofs for the ✓ z () > & ∃ and bounded ∀.
x C y ✓ z () x ✓ z ^ y 2 z

✤

The proof in the HF x = y () needs^ y ✓ x 450 lines.
calculus x ✓ y under

The point of all this is that (??) and (??) can be proved by a simultaneous
✤ Fills a major gap in various presentations, including S!wierczkowski’s.
induction:
QR(x, x0 ), QR(y, y 0 ) ` (x 2 y ! Pf bx0 2 y 0 c{x0 ,y0 } ) ^ (x ✓ y ! Pf bx0 ✓ y 0 c{x0 ,y0 } )
=) Pf bx0 ✓ y 0 c{x0 ,y0 } ^ Pf bx0 2 y 0 c{x0 ,y0 }
1
2
1
2
() Pf bx0 C x0 ✓ y 0 c{x0 ,x0 ,y0 }
1
2
1 2

theorem Goedel II:
assumes Con: " ¬ {} ` Fls"
shows
" ¬ {} ` Neg (PfP pFls q)"
proof from Con Goedel I obtain
where diag: "{} `
IFF Neg (PfP p q)" " ¬ {} ` "
and gnd: "ground fm "
by metis
have "{PfP p q} ` PfP pPfP p qq"
by (auto simp: Provability ground fm aux def supp conv fresh)
moreover have "{PfP p q} ` PfP pNeg (PfP p q) q"
apply (rule MonPon PfP implies PfP [OF gnd])
apply (auto simp: ground fm aux def supp conv fresh) using diag
by (metis Assume ContraProve Iff MP left Iff MP left’ Neg Neg iff)
moreover have "ground fm (PfP p q)"
by (auto simp: ground fm aux def supp conv fresh)
ultimately have "{PfP p q} ` PfP pFls q" using PfP quot contra
by (metis (no types) anti deduction cut2)
thus " ¬ {} ` Neg (PfP pFls q)"
by (metis Iff MP2 same Neg mono cut1 diag)
qed
What Did We Learn?
✤

Some highly compressed proofs were finally made explicit.

✤

The entire proof development can be examined interactively.

✤

The nominal package can cope with very large developments…
(BUT: performance issues, some repetitive notation, complications in
accepting function definitions)

✤

<9 months for the first theorem, a further 4 for the second

✤

Under 16 000 lines of proof script in all.
Conclusions
✤

the first-ever machine formalisation of Gödel’s second incompleteness
theorem

✤

using both nominal and de Bruijn syntax for bound variables

✤

within an axiom system for hereditarily finite set theory

✤

conducted using Isabelle/HOL.
Many thanks to Christian Urban for help at critical points!
Also Jesse Alama, Dana Scott.

Más contenido relacionado

La actualidad más candente

Discrete Mathematical Structures - Fundamentals of Logic - Principle of duality
Discrete Mathematical Structures - Fundamentals of Logic - Principle of dualityDiscrete Mathematical Structures - Fundamentals of Logic - Principle of duality
Discrete Mathematical Structures - Fundamentals of Logic - Principle of dualityLakshmi R
 
Fractional Calculus
Fractional CalculusFractional Calculus
Fractional CalculusVRRITC
 
X2 T01 07 nth roots of unity
X2 T01 07 nth roots of unityX2 T01 07 nth roots of unity
X2 T01 07 nth roots of unityNigel Simmons
 
Laplace transforms
Laplace transforms Laplace transforms
Laplace transforms yash patel
 
제4강 명제와 논리-정보
제4강 명제와 논리-정보제4강 명제와 논리-정보
제4강 명제와 논리-정보csungwoo
 
Lie-Trotter-Suzuki分解、特にフラクタル分解について
Lie-Trotter-Suzuki分解、特にフラクタル分解についてLie-Trotter-Suzuki分解、特にフラクタル分解について
Lie-Trotter-Suzuki分解、特にフラクタル分解についてMaho Nakata
 
लम्ब वृत्तीय बेलन
लम्ब वृत्तीय बेलन लम्ब वृत्तीय बेलन
लम्ब वृत्तीय बेलन ramniwas shukla
 
Policy Gradient Theorem
Policy Gradient TheoremPolicy Gradient Theorem
Policy Gradient TheoremAshwin Rao
 
ガウス過程入門
ガウス過程入門ガウス過程入門
ガウス過程入門ShoShimoyama
 
1575 numerical differentiation and integration
1575 numerical differentiation and integration1575 numerical differentiation and integration
1575 numerical differentiation and integrationDr Fereidoun Dejahang
 

La actualidad más candente (14)

Discrete Mathematical Structures - Fundamentals of Logic - Principle of duality
Discrete Mathematical Structures - Fundamentals of Logic - Principle of dualityDiscrete Mathematical Structures - Fundamentals of Logic - Principle of duality
Discrete Mathematical Structures - Fundamentals of Logic - Principle of duality
 
Bethe salpeter equation
Bethe salpeter equationBethe salpeter equation
Bethe salpeter equation
 
Fractional Calculus
Fractional CalculusFractional Calculus
Fractional Calculus
 
X2 T01 07 nth roots of unity
X2 T01 07 nth roots of unityX2 T01 07 nth roots of unity
X2 T01 07 nth roots of unity
 
Laplace transforms
Laplace transforms Laplace transforms
Laplace transforms
 
제4강 명제와 논리-정보
제4강 명제와 논리-정보제4강 명제와 논리-정보
제4강 명제와 논리-정보
 
Lie-Trotter-Suzuki分解、特にフラクタル分解について
Lie-Trotter-Suzuki分解、特にフラクタル分解についてLie-Trotter-Suzuki分解、特にフラクタル分解について
Lie-Trotter-Suzuki分解、特にフラクタル分解について
 
Runge Kutta Method
Runge Kutta MethodRunge Kutta Method
Runge Kutta Method
 
Line integral.ppt
Line integral.pptLine integral.ppt
Line integral.ppt
 
लम्ब वृत्तीय बेलन
लम्ब वृत्तीय बेलन लम्ब वृत्तीय बेलन
लम्ब वृत्तीय बेलन
 
Policy Gradient Theorem
Policy Gradient TheoremPolicy Gradient Theorem
Policy Gradient Theorem
 
ガウス過程入門
ガウス過程入門ガウス過程入門
ガウス過程入門
 
FPDE presentation
FPDE presentationFPDE presentation
FPDE presentation
 
1575 numerical differentiation and integration
1575 numerical differentiation and integration1575 numerical differentiation and integration
1575 numerical differentiation and integration
 

Destacado

GEB Gödel, Escher, Bach: An Eternal Golden Braid
GEB Gödel, Escher, Bach: An Eternal Golden BraidGEB Gödel, Escher, Bach: An Eternal Golden Braid
GEB Gödel, Escher, Bach: An Eternal Golden Braidgordana comic
 
Incompleteness without Godel Numberings
Incompleteness without Godel NumberingsIncompleteness without Godel Numberings
Incompleteness without Godel NumberingsSatvik Beri
 
Godels First Incompleteness Theorem
Godels First Incompleteness TheoremGodels First Incompleteness Theorem
Godels First Incompleteness Theoremmmanning02474
 
Incompleteness Theorems: Logical Necessity of Inconsistency
Incompleteness Theorems:  Logical Necessity of InconsistencyIncompleteness Theorems:  Logical Necessity of Inconsistency
Incompleteness Theorems: Logical Necessity of InconsistencyCarl Hewitt
 
The incompleteness of reason
The incompleteness of reasonThe incompleteness of reason
The incompleteness of reasonSubhayan Mukerjee
 

Destacado (6)

GEB Gödel, Escher, Bach: An Eternal Golden Braid
GEB Gödel, Escher, Bach: An Eternal Golden BraidGEB Gödel, Escher, Bach: An Eternal Golden Braid
GEB Gödel, Escher, Bach: An Eternal Golden Braid
 
Computability
Computability Computability
Computability
 
Incompleteness without Godel Numberings
Incompleteness without Godel NumberingsIncompleteness without Godel Numberings
Incompleteness without Godel Numberings
 
Godels First Incompleteness Theorem
Godels First Incompleteness TheoremGodels First Incompleteness Theorem
Godels First Incompleteness Theorem
 
Incompleteness Theorems: Logical Necessity of Inconsistency
Incompleteness Theorems:  Logical Necessity of InconsistencyIncompleteness Theorems:  Logical Necessity of Inconsistency
Incompleteness Theorems: Logical Necessity of Inconsistency
 
The incompleteness of reason
The incompleteness of reasonThe incompleteness of reason
The incompleteness of reason
 

Similar a A Machine-Assisted Proof of Gödel's Incompleteness Theorems

Cs229 notes8
Cs229 notes8Cs229 notes8
Cs229 notes8VuTran231
 
Applications of partial differentiation
Applications of partial differentiationApplications of partial differentiation
Applications of partial differentiationVaibhav Tandel
 
2015 CMS Winter Meeting Poster
2015 CMS Winter Meeting Poster2015 CMS Winter Meeting Poster
2015 CMS Winter Meeting PosterChelsea Battell
 
MetiTarski: An Automatic Prover for Real-Valued Special Functions
MetiTarski: An Automatic Prover for Real-Valued Special FunctionsMetiTarski: An Automatic Prover for Real-Valued Special Functions
MetiTarski: An Automatic Prover for Real-Valued Special FunctionsLawrence Paulson
 
The Chase in Database Theory
The Chase in Database TheoryThe Chase in Database Theory
The Chase in Database TheoryJan Hidders
 
Sentient Arithmetic and Godel's Incompleteness Theorems
Sentient Arithmetic and Godel's Incompleteness TheoremsSentient Arithmetic and Godel's Incompleteness Theorems
Sentient Arithmetic and Godel's Incompleteness TheoremsKannan Nambiar
 
Csr2011 june17 14_00_bulatov
Csr2011 june17 14_00_bulatovCsr2011 june17 14_00_bulatov
Csr2011 june17 14_00_bulatovCSR2011
 
Csr2011 june17 14_00_bulatov
Csr2011 june17 14_00_bulatovCsr2011 june17 14_00_bulatov
Csr2011 june17 14_00_bulatovCSR2011
 
Du Calcul des prédicats vers Prolog
Du Calcul des prédicats vers PrologDu Calcul des prédicats vers Prolog
Du Calcul des prédicats vers PrologSerge Garlatti
 
MVT mean value theorem نظرية القيمة المتوسطة
MVT mean value theorem نظرية القيمة المتوسطةMVT mean value theorem نظرية القيمة المتوسطة
MVT mean value theorem نظرية القيمة المتوسطةDr. Karrar Alwash
 
Boolean Programs and Quantified Propositional Proof System -
Boolean Programs and Quantified Propositional Proof System - Boolean Programs and Quantified Propositional Proof System -
Boolean Programs and Quantified Propositional Proof System - Michael Soltys
 
Jensen's inequality, EM 알고리즘
Jensen's inequality, EM 알고리즘 Jensen's inequality, EM 알고리즘
Jensen's inequality, EM 알고리즘 Jungkyu Lee
 
Machine learning (9)
Machine learning (9)Machine learning (9)
Machine learning (9)NYversity
 
Machine learning (1)
Machine learning (1)Machine learning (1)
Machine learning (1)NYversity
 
a logical approach to abstract algebra (apuntes).pdf
a logical approach to abstract algebra (apuntes).pdfa logical approach to abstract algebra (apuntes).pdf
a logical approach to abstract algebra (apuntes).pdfcibeyo cibeyo
 
5.1 anti derivatives
5.1 anti derivatives5.1 anti derivatives
5.1 anti derivativesmath265
 
Applicationofpartialderivativeswithtwovariables 140225070102-phpapp01 (1)
Applicationofpartialderivativeswithtwovariables 140225070102-phpapp01 (1)Applicationofpartialderivativeswithtwovariables 140225070102-phpapp01 (1)
Applicationofpartialderivativeswithtwovariables 140225070102-phpapp01 (1)shreemadghodasra
 

Similar a A Machine-Assisted Proof of Gödel's Incompleteness Theorems (20)

Cs229 notes8
Cs229 notes8Cs229 notes8
Cs229 notes8
 
Applications of partial differentiation
Applications of partial differentiationApplications of partial differentiation
Applications of partial differentiation
 
2015 CMS Winter Meeting Poster
2015 CMS Winter Meeting Poster2015 CMS Winter Meeting Poster
2015 CMS Winter Meeting Poster
 
MetiTarski: An Automatic Prover for Real-Valued Special Functions
MetiTarski: An Automatic Prover for Real-Valued Special FunctionsMetiTarski: An Automatic Prover for Real-Valued Special Functions
MetiTarski: An Automatic Prover for Real-Valued Special Functions
 
The Chase in Database Theory
The Chase in Database TheoryThe Chase in Database Theory
The Chase in Database Theory
 
Sentient Arithmetic and Godel's Incompleteness Theorems
Sentient Arithmetic and Godel's Incompleteness TheoremsSentient Arithmetic and Godel's Incompleteness Theorems
Sentient Arithmetic and Godel's Incompleteness Theorems
 
Csr2011 june17 14_00_bulatov
Csr2011 june17 14_00_bulatovCsr2011 june17 14_00_bulatov
Csr2011 june17 14_00_bulatov
 
Csr2011 june17 14_00_bulatov
Csr2011 june17 14_00_bulatovCsr2011 june17 14_00_bulatov
Csr2011 june17 14_00_bulatov
 
C2.0 propositional logic
C2.0 propositional logicC2.0 propositional logic
C2.0 propositional logic
 
Du Calcul des prédicats vers Prolog
Du Calcul des prédicats vers PrologDu Calcul des prédicats vers Prolog
Du Calcul des prédicats vers Prolog
 
Predicate Calculus
Predicate CalculusPredicate Calculus
Predicate Calculus
 
MVT mean value theorem نظرية القيمة المتوسطة
MVT mean value theorem نظرية القيمة المتوسطةMVT mean value theorem نظرية القيمة المتوسطة
MVT mean value theorem نظرية القيمة المتوسطة
 
Boolean Programs and Quantified Propositional Proof System -
Boolean Programs and Quantified Propositional Proof System - Boolean Programs and Quantified Propositional Proof System -
Boolean Programs and Quantified Propositional Proof System -
 
Jensen's inequality, EM 알고리즘
Jensen's inequality, EM 알고리즘 Jensen's inequality, EM 알고리즘
Jensen's inequality, EM 알고리즘
 
Machine learning (9)
Machine learning (9)Machine learning (9)
Machine learning (9)
 
Machine learning (1)
Machine learning (1)Machine learning (1)
Machine learning (1)
 
a logical approach to abstract algebra (apuntes).pdf
a logical approach to abstract algebra (apuntes).pdfa logical approach to abstract algebra (apuntes).pdf
a logical approach to abstract algebra (apuntes).pdf
 
Pl vol1
Pl vol1Pl vol1
Pl vol1
 
5.1 anti derivatives
5.1 anti derivatives5.1 anti derivatives
5.1 anti derivatives
 
Applicationofpartialderivativeswithtwovariables 140225070102-phpapp01 (1)
Applicationofpartialderivativeswithtwovariables 140225070102-phpapp01 (1)Applicationofpartialderivativeswithtwovariables 140225070102-phpapp01 (1)
Applicationofpartialderivativeswithtwovariables 140225070102-phpapp01 (1)
 

Más de Lawrence Paulson

Proving Security Protocols Correct
Proving Security Protocols CorrectProving Security Protocols Correct
Proving Security Protocols CorrectLawrence Paulson
 
MetiTarski's menagerie of cooperating systems
MetiTarski's menagerie of cooperating systemsMetiTarski's menagerie of cooperating systems
MetiTarski's menagerie of cooperating systemsLawrence Paulson
 
Automated theorem proving for special functions: the next phase
Automated theorem proving for special functions: the next phaseAutomated theorem proving for special functions: the next phase
Automated theorem proving for special functions: the next phaseLawrence Paulson
 
Theorem proving and the real numbers: overview and challenges
Theorem proving and the real numbers: overview and challengesTheorem proving and the real numbers: overview and challenges
Theorem proving and the real numbers: overview and challengesLawrence Paulson
 
Source-Level Proof Reconstruction for Interactive Proving
Source-Level Proof Reconstruction for Interactive ProvingSource-Level Proof Reconstruction for Interactive Proving
Source-Level Proof Reconstruction for Interactive ProvingLawrence Paulson
 
Defining Functions on Equivalence Classes
Defining Functions on Equivalence ClassesDefining Functions on Equivalence Classes
Defining Functions on Equivalence ClassesLawrence Paulson
 
Organizing Numerical Theories using Axiomatic Type Classes
Organizing Numerical Theories using Axiomatic Type ClassesOrganizing Numerical Theories using Axiomatic Type Classes
Organizing Numerical Theories using Axiomatic Type ClassesLawrence Paulson
 
A Generic Tableau Prover and Its Integration with Isabelle
A Generic Tableau Prover and Its Integration with IsabelleA Generic Tableau Prover and Its Integration with Isabelle
A Generic Tableau Prover and Its Integration with IsabelleLawrence Paulson
 
Mechanized Proofs for a Recursive Authentication Protocol
Mechanized Proofs for a Recursive Authentication ProtocolMechanized Proofs for a Recursive Authentication Protocol
Mechanized Proofs for a Recursive Authentication ProtocolLawrence Paulson
 
Mechanizing set theory: cardinal arithmetic and the axiom of choice
Mechanizing set theory: cardinal arithmetic and the axiom of choiceMechanizing set theory: cardinal arithmetic and the axiom of choice
Mechanizing set theory: cardinal arithmetic and the axiom of choiceLawrence Paulson
 
The Relative Consistency of the Axiom of Choice — Mechanized Using Isabelle/ZF
The Relative Consistency of the Axiom of Choice — Mechanized Using Isabelle/ZFThe Relative Consistency of the Axiom of Choice — Mechanized Using Isabelle/ZF
The Relative Consistency of the Axiom of Choice — Mechanized Using Isabelle/ZFLawrence Paulson
 
The Reflection Theorem: Formalizing Meta-Theoretic Reasoning
The Reflection Theorem: Formalizing Meta-Theoretic ReasoningThe Reflection Theorem: Formalizing Meta-Theoretic Reasoning
The Reflection Theorem: Formalizing Meta-Theoretic ReasoningLawrence Paulson
 
Proving Properties of Security Protocols by Induction
Proving Properties of Security Protocols by InductionProving Properties of Security Protocols by Induction
Proving Properties of Security Protocols by InductionLawrence Paulson
 

Más de Lawrence Paulson (13)

Proving Security Protocols Correct
Proving Security Protocols CorrectProving Security Protocols Correct
Proving Security Protocols Correct
 
MetiTarski's menagerie of cooperating systems
MetiTarski's menagerie of cooperating systemsMetiTarski's menagerie of cooperating systems
MetiTarski's menagerie of cooperating systems
 
Automated theorem proving for special functions: the next phase
Automated theorem proving for special functions: the next phaseAutomated theorem proving for special functions: the next phase
Automated theorem proving for special functions: the next phase
 
Theorem proving and the real numbers: overview and challenges
Theorem proving and the real numbers: overview and challengesTheorem proving and the real numbers: overview and challenges
Theorem proving and the real numbers: overview and challenges
 
Source-Level Proof Reconstruction for Interactive Proving
Source-Level Proof Reconstruction for Interactive ProvingSource-Level Proof Reconstruction for Interactive Proving
Source-Level Proof Reconstruction for Interactive Proving
 
Defining Functions on Equivalence Classes
Defining Functions on Equivalence ClassesDefining Functions on Equivalence Classes
Defining Functions on Equivalence Classes
 
Organizing Numerical Theories using Axiomatic Type Classes
Organizing Numerical Theories using Axiomatic Type ClassesOrganizing Numerical Theories using Axiomatic Type Classes
Organizing Numerical Theories using Axiomatic Type Classes
 
A Generic Tableau Prover and Its Integration with Isabelle
A Generic Tableau Prover and Its Integration with IsabelleA Generic Tableau Prover and Its Integration with Isabelle
A Generic Tableau Prover and Its Integration with Isabelle
 
Mechanized Proofs for a Recursive Authentication Protocol
Mechanized Proofs for a Recursive Authentication ProtocolMechanized Proofs for a Recursive Authentication Protocol
Mechanized Proofs for a Recursive Authentication Protocol
 
Mechanizing set theory: cardinal arithmetic and the axiom of choice
Mechanizing set theory: cardinal arithmetic and the axiom of choiceMechanizing set theory: cardinal arithmetic and the axiom of choice
Mechanizing set theory: cardinal arithmetic and the axiom of choice
 
The Relative Consistency of the Axiom of Choice — Mechanized Using Isabelle/ZF
The Relative Consistency of the Axiom of Choice — Mechanized Using Isabelle/ZFThe Relative Consistency of the Axiom of Choice — Mechanized Using Isabelle/ZF
The Relative Consistency of the Axiom of Choice — Mechanized Using Isabelle/ZF
 
The Reflection Theorem: Formalizing Meta-Theoretic Reasoning
The Reflection Theorem: Formalizing Meta-Theoretic ReasoningThe Reflection Theorem: Formalizing Meta-Theoretic Reasoning
The Reflection Theorem: Formalizing Meta-Theoretic Reasoning
 
Proving Properties of Security Protocols by Induction
Proving Properties of Security Protocols by InductionProving Properties of Security Protocols by Induction
Proving Properties of Security Protocols by Induction
 

Último

Cloud Frontiers: A Deep Dive into Serverless Spatial Data and FME
Cloud Frontiers:  A Deep Dive into Serverless Spatial Data and FMECloud Frontiers:  A Deep Dive into Serverless Spatial Data and FME
Cloud Frontiers: A Deep Dive into Serverless Spatial Data and FMESafe Software
 
AXA XL - Insurer Innovation Award Americas 2024
AXA XL - Insurer Innovation Award Americas 2024AXA XL - Insurer Innovation Award Americas 2024
AXA XL - Insurer Innovation Award Americas 2024The Digital Insurer
 
Why Teams call analytics are critical to your entire business
Why Teams call analytics are critical to your entire businessWhy Teams call analytics are critical to your entire business
Why Teams call analytics are critical to your entire businesspanagenda
 
Apidays New York 2024 - Scaling API-first by Ian Reasor and Radu Cotescu, Adobe
Apidays New York 2024 - Scaling API-first by Ian Reasor and Radu Cotescu, AdobeApidays New York 2024 - Scaling API-first by Ian Reasor and Radu Cotescu, Adobe
Apidays New York 2024 - Scaling API-first by Ian Reasor and Radu Cotescu, Adobeapidays
 
Powerful Google developer tools for immediate impact! (2023-24 C)
Powerful Google developer tools for immediate impact! (2023-24 C)Powerful Google developer tools for immediate impact! (2023-24 C)
Powerful Google developer tools for immediate impact! (2023-24 C)wesley chun
 
Ransomware_Q4_2023. The report. [EN].pdf
Ransomware_Q4_2023. The report. [EN].pdfRansomware_Q4_2023. The report. [EN].pdf
Ransomware_Q4_2023. The report. [EN].pdfOverkill Security
 
Apidays Singapore 2024 - Building Digital Trust in a Digital Economy by Veron...
Apidays Singapore 2024 - Building Digital Trust in a Digital Economy by Veron...Apidays Singapore 2024 - Building Digital Trust in a Digital Economy by Veron...
Apidays Singapore 2024 - Building Digital Trust in a Digital Economy by Veron...apidays
 
Axa Assurance Maroc - Insurer Innovation Award 2024
Axa Assurance Maroc - Insurer Innovation Award 2024Axa Assurance Maroc - Insurer Innovation Award 2024
Axa Assurance Maroc - Insurer Innovation Award 2024The Digital Insurer
 
How to Troubleshoot Apps for the Modern Connected Worker
How to Troubleshoot Apps for the Modern Connected WorkerHow to Troubleshoot Apps for the Modern Connected Worker
How to Troubleshoot Apps for the Modern Connected WorkerThousandEyes
 
MS Copilot expands with MS Graph connectors
MS Copilot expands with MS Graph connectorsMS Copilot expands with MS Graph connectors
MS Copilot expands with MS Graph connectorsNanddeep Nachan
 
Data Cloud, More than a CDP by Matt Robison
Data Cloud, More than a CDP by Matt RobisonData Cloud, More than a CDP by Matt Robison
Data Cloud, More than a CDP by Matt RobisonAnna Loughnan Colquhoun
 
ProductAnonymous-April2024-WinProductDiscovery-MelissaKlemke
ProductAnonymous-April2024-WinProductDiscovery-MelissaKlemkeProductAnonymous-April2024-WinProductDiscovery-MelissaKlemke
ProductAnonymous-April2024-WinProductDiscovery-MelissaKlemkeProduct Anonymous
 
Repurposing LNG terminals for Hydrogen Ammonia: Feasibility and Cost Saving
Repurposing LNG terminals for Hydrogen Ammonia: Feasibility and Cost SavingRepurposing LNG terminals for Hydrogen Ammonia: Feasibility and Cost Saving
Repurposing LNG terminals for Hydrogen Ammonia: Feasibility and Cost SavingEdi Saputra
 
Connector Corner: Accelerate revenue generation using UiPath API-centric busi...
Connector Corner: Accelerate revenue generation using UiPath API-centric busi...Connector Corner: Accelerate revenue generation using UiPath API-centric busi...
Connector Corner: Accelerate revenue generation using UiPath API-centric busi...DianaGray10
 
ICT role in 21st century education and its challenges
ICT role in 21st century education and its challengesICT role in 21st century education and its challenges
ICT role in 21st century education and its challengesrafiqahmad00786416
 
Real Time Object Detection Using Open CV
Real Time Object Detection Using Open CVReal Time Object Detection Using Open CV
Real Time Object Detection Using Open CVKhem
 
presentation ICT roal in 21st century education
presentation ICT roal in 21st century educationpresentation ICT roal in 21st century education
presentation ICT roal in 21st century educationjfdjdjcjdnsjd
 
Strategies for Landing an Oracle DBA Job as a Fresher
Strategies for Landing an Oracle DBA Job as a FresherStrategies for Landing an Oracle DBA Job as a Fresher
Strategies for Landing an Oracle DBA Job as a FresherRemote DBA Services
 
Emergent Methods: Multi-lingual narrative tracking in the news - real-time ex...
Emergent Methods: Multi-lingual narrative tracking in the news - real-time ex...Emergent Methods: Multi-lingual narrative tracking in the news - real-time ex...
Emergent Methods: Multi-lingual narrative tracking in the news - real-time ex...Zilliz
 
Polkadot JAM Slides - Token2049 - By Dr. Gavin Wood
Polkadot JAM Slides - Token2049 - By Dr. Gavin WoodPolkadot JAM Slides - Token2049 - By Dr. Gavin Wood
Polkadot JAM Slides - Token2049 - By Dr. Gavin WoodJuan lago vázquez
 

Último (20)

Cloud Frontiers: A Deep Dive into Serverless Spatial Data and FME
Cloud Frontiers:  A Deep Dive into Serverless Spatial Data and FMECloud Frontiers:  A Deep Dive into Serverless Spatial Data and FME
Cloud Frontiers: A Deep Dive into Serverless Spatial Data and FME
 
AXA XL - Insurer Innovation Award Americas 2024
AXA XL - Insurer Innovation Award Americas 2024AXA XL - Insurer Innovation Award Americas 2024
AXA XL - Insurer Innovation Award Americas 2024
 
Why Teams call analytics are critical to your entire business
Why Teams call analytics are critical to your entire businessWhy Teams call analytics are critical to your entire business
Why Teams call analytics are critical to your entire business
 
Apidays New York 2024 - Scaling API-first by Ian Reasor and Radu Cotescu, Adobe
Apidays New York 2024 - Scaling API-first by Ian Reasor and Radu Cotescu, AdobeApidays New York 2024 - Scaling API-first by Ian Reasor and Radu Cotescu, Adobe
Apidays New York 2024 - Scaling API-first by Ian Reasor and Radu Cotescu, Adobe
 
Powerful Google developer tools for immediate impact! (2023-24 C)
Powerful Google developer tools for immediate impact! (2023-24 C)Powerful Google developer tools for immediate impact! (2023-24 C)
Powerful Google developer tools for immediate impact! (2023-24 C)
 
Ransomware_Q4_2023. The report. [EN].pdf
Ransomware_Q4_2023. The report. [EN].pdfRansomware_Q4_2023. The report. [EN].pdf
Ransomware_Q4_2023. The report. [EN].pdf
 
Apidays Singapore 2024 - Building Digital Trust in a Digital Economy by Veron...
Apidays Singapore 2024 - Building Digital Trust in a Digital Economy by Veron...Apidays Singapore 2024 - Building Digital Trust in a Digital Economy by Veron...
Apidays Singapore 2024 - Building Digital Trust in a Digital Economy by Veron...
 
Axa Assurance Maroc - Insurer Innovation Award 2024
Axa Assurance Maroc - Insurer Innovation Award 2024Axa Assurance Maroc - Insurer Innovation Award 2024
Axa Assurance Maroc - Insurer Innovation Award 2024
 
How to Troubleshoot Apps for the Modern Connected Worker
How to Troubleshoot Apps for the Modern Connected WorkerHow to Troubleshoot Apps for the Modern Connected Worker
How to Troubleshoot Apps for the Modern Connected Worker
 
MS Copilot expands with MS Graph connectors
MS Copilot expands with MS Graph connectorsMS Copilot expands with MS Graph connectors
MS Copilot expands with MS Graph connectors
 
Data Cloud, More than a CDP by Matt Robison
Data Cloud, More than a CDP by Matt RobisonData Cloud, More than a CDP by Matt Robison
Data Cloud, More than a CDP by Matt Robison
 
ProductAnonymous-April2024-WinProductDiscovery-MelissaKlemke
ProductAnonymous-April2024-WinProductDiscovery-MelissaKlemkeProductAnonymous-April2024-WinProductDiscovery-MelissaKlemke
ProductAnonymous-April2024-WinProductDiscovery-MelissaKlemke
 
Repurposing LNG terminals for Hydrogen Ammonia: Feasibility and Cost Saving
Repurposing LNG terminals for Hydrogen Ammonia: Feasibility and Cost SavingRepurposing LNG terminals for Hydrogen Ammonia: Feasibility and Cost Saving
Repurposing LNG terminals for Hydrogen Ammonia: Feasibility and Cost Saving
 
Connector Corner: Accelerate revenue generation using UiPath API-centric busi...
Connector Corner: Accelerate revenue generation using UiPath API-centric busi...Connector Corner: Accelerate revenue generation using UiPath API-centric busi...
Connector Corner: Accelerate revenue generation using UiPath API-centric busi...
 
ICT role in 21st century education and its challenges
ICT role in 21st century education and its challengesICT role in 21st century education and its challenges
ICT role in 21st century education and its challenges
 
Real Time Object Detection Using Open CV
Real Time Object Detection Using Open CVReal Time Object Detection Using Open CV
Real Time Object Detection Using Open CV
 
presentation ICT roal in 21st century education
presentation ICT roal in 21st century educationpresentation ICT roal in 21st century education
presentation ICT roal in 21st century education
 
Strategies for Landing an Oracle DBA Job as a Fresher
Strategies for Landing an Oracle DBA Job as a FresherStrategies for Landing an Oracle DBA Job as a Fresher
Strategies for Landing an Oracle DBA Job as a Fresher
 
Emergent Methods: Multi-lingual narrative tracking in the news - real-time ex...
Emergent Methods: Multi-lingual narrative tracking in the news - real-time ex...Emergent Methods: Multi-lingual narrative tracking in the news - real-time ex...
Emergent Methods: Multi-lingual narrative tracking in the news - real-time ex...
 
Polkadot JAM Slides - Token2049 - By Dr. Gavin Wood
Polkadot JAM Slides - Token2049 - By Dr. Gavin WoodPolkadot JAM Slides - Token2049 - By Dr. Gavin Wood
Polkadot JAM Slides - Token2049 - By Dr. Gavin Wood
 

A Machine-Assisted Proof of Gödel's Incompleteness Theorems

  • 1. A Machine-Assisted Proof of Gödel's Incompleteness Theorems Lawrence C. Paulson, Computer Laboratory, University of Cambridge
  • 2. The most misunderstood theorems in mathematics ✤ Gödel’s theorems have highly technical, syntactic proofs. 1. Every “strong enough” formal system is incomplete, in that at least one formula can neither be proved nor disproved. 2. And if such a formal system admits a proof of its own consistency, then it is actually inconsistent. ✤ For the first time, both of Gödel’s proofs have been mechanised, following a paper by S!wierczkowski (2003) ✤ The machine proof, in the structured Isar language, is complete, almost readable, and can be perused interactively.
  • 3. Hereditarily finite set theory ✤ A hereditarily finite set is a finite set of HF sets. ✤ Many mathematical constructions, including natural numbers and sequences, can be defined as in standard set theory. ✤ HF set theory is equivalent to Peano Arithmetic via the mapping f(x) = X f(y) {2 | y 2 x}
  • 4. Benefits of Using HF Set Theory ✤ Can use standard definitions of pairing and sequences. ✤ The first incompleteness theorem requires an HF development of the natural numbers, induction, etc., but not addition. ✤ The second incompleteness theorem requires operations on sequences and addition, but not multiplication. ✤ No need for least common multiples, prime numbers or the Chinese remainder theorem.
  • 5. The Axioms of HF Set Theory ✤ 0 denotes the empty set ✤ x ◁ y denotes the set x extended with the element y. ✤ There are no other function symbols. ✤ Union, intersection, etc can be shown to exist by induction.
  • 6. Stages of the Proofs ✤ The syntax of a first-order theory is formalised: terms, formulas, substitution... ✤ A system of coding to formalise the calculus within itself. The code of # is a term, written ⌜#⌝. ✤ A deductive calculus for sequents of the form " ⊦ # (typically for Peano arithmetic, but here HF) ✤ Syntactic predicates to recognise codes of terms, substitution, axioms, etc. ✤ Meta-theory to relate truth and provability. E.g. “all true $ formulas are theorems”. $ formulas are built using % & ∃ and bounded ∀. ✤ Finally the predicate Pf, such that ⊦ # ⟺ ⊦ Pf ⌜#⌝.
  • 7. First Incompleteness Theorem ✤ To prove Gödel’s first incompleteness theorem, construct ' that expresses that ' is not provable. ✤ It follows (provided the calculus is consistent) that neither ' nor its negation can be proved. ✤ Need to show that substitution behaves like a function. ✤ Requires a detailed proof in the calculus, ✤ … alternatively, other detailed calculations.
  • 8. Second Incompleteness Theorem ✤ This crucial lemma for Gödel’s second incompleteness theorem is proved by induction over the construction of # as a $ formula. ✤ It requires generalising the statement above to allow the formula # to contain free variables. ✤ complex technicalities ✤ lengthy deductions in the calculus
  • 9. Proving Theorems in the Calculus ✤ Gödel knew that formal proofs were difficult.They could be eliminated, but at what cost? ✤ By coding all predicates as executable functions, and proving a meta-theorem, Gödel reduced provability to truth. ✤ But then only bounded quantifiers can be used, with tricky arithmetical proofs that the bounds are adequate. ✤ With $ formulas, provability is reduced to truth for most desired properties, with no tricky proofs about bounds. ✤ Instead, some straightforward inductions need to be formalised in the calculus. ✤ The second theorem requires working in the calculus anyway.
  • 10. Isabelle/HOL and Nominal ✤ a proof assistant for higherorder logic ✤ Free names are significant, but not once they are bound. ✤ much automation to hide the underlying proof calculus ✤ ✤ support for recursive functions and inductive sets Syntax involving variable binding can be defined using recursion, provided variables are used “sensibly”. ✤ During proof by induction, bound variable names can be guaranteed not to clash with specified other terms. ✤ the nominal package, for working with named variables
  • 11. De Bruijn Indexes ✤ This approach to variable binding replaces names by numbers. ✤ 0 denotes the innermost bound variable, 1 for the next, etc. ✤ This approach destroys readability, but substitution and abstraction are very easy to define. ✤ During coding, formulas are translated into the de Bruijn format. ✤ And so there is no need to formalise the nominal theory within HF.
  • 12. 1 The Isabelle/HOL proof development: fundamentals. Defining Terms and Formulas z = 0 $ 8x [x 62 z] z = x C y $ 8u [u 2 z $ u 2 x _ u = y] (0) ^ 8xy [ (x) ^ (y) ! (x C y)] ! 8x [ (x)] (HF1) (HF2) (HF3) nominal datatype tm = Zero | Var name | Eats tm tm nominal datatype fm = Mem tm tm ( infixr "IN" | Eq tm tm ( infixr "EQ" | Disj fm fm ( infixr "OR" | Neg fm | Ex x::name f::fm binds x 150) 150) 130) in f The HF calculus includes an existential quantifier, denoted Ex, which involves variable binding via the nominal framework. The infixr declarations provide an alternative syntax for the membership relation, the equality relation, and disjunction. A formula can also be a negation. The other logical connectives Variable binding are introduced later as abbreviations. formalised to formalise, but here it is straightSubstitution is often problematicalusing nominal forward. Substitution of a term x for a variable i is defined as follows: nominal primrec subst :: "name ) tm ) tm ) tm" where "subst i x Zero = Zero"
  • 13. The HF calculus includes an existential quantifier, denoted Ex, which involves variable binding via the nominal framework. The infixr declarations provide an alternative syntax for the membership relation, the equality relation, and disjunction. A formula can also be a negation. The other logical connectives are introduced later as abbreviations. Substitution is often problematical to formalise, but here it is straightforward. Substitution of a term x for a variable i is defined as follows: Defining Substitution nominal primrec subst :: "name ) tm ) tm ) tm" where "subst i x Zero = Zero" For substitution within a formula, we normally expect issues concerning | "subst i x (Var k) = (if i=k then x else Var k)" the "subst iofx a(Eats t variable. Note that thet) (subst substituting the bound u) = Eats (subst i x result of i x u)" | capture term x for the variable i in the formula A is written A(i::=x). 1 nominal primrec subst fm :: "fm ) name ) tm ) fm" where Mem: "(Mem t u)(i::=x) = Mem (subst i x t) (subst i x u)" | Eq: "(Eq t u)(i::=x) = Eq (subst i x t) (subst i x u)" | Disj: "(Disj A B)(i::=x) = Disj (A(i::=x)) (B(i::=x))" | Neg: "(Neg A)(i::=x) = Neg (A(i::=x))" | Ex: "atom j ] (i, x) =) (Ex j A)(i::=x) = Ex j (A(i::=x))" Substitution is again straightforward in the first four cases (membership, Properties of equality, disjunction, negation). be the existential case, the precondition In The variable j must atom j ] (i, x) (pronounced “j is fresh for i and xsubstitution have ”) essentially says that fresh di↵erent names i and j must be for i and x with j not free in x. We do not need to simple proofs. supply a mechanism for renaming the bound variable, as that is part of the nominal framework, which in most cases will choose a su ciently fresh bound
  • 14. supply a mechanism for renaming the bound variable, as that is part of the nominal framework, which in most cases will choose a su ciently fresh bound variable at the outset. The usual properties of substitution (commutativity, for example) have simple proofs by induction on formulas. In contrast, ?) needed to combine three substitution lemmas in a simultaneous proof by induction, a delicate argument involving 1900 lines of Coq. The HF proof system is an inductively defined predicate, where H ` A means that the formula A is provable from the set of formulas H . Defining the HF Calculus inductive hfthm :: "fm set ) fm ) bool" ( infixl " `" 55) where Hyp: "A 2 H =) H ` A" | Extra: "H ` extra axiom" | Bool: "A 2 boolean axioms =) H ` A" | Eq: "A 2 equality axioms =) H ` A" | Spec: "A 2 special axioms =) H ` A" | HF: "A 2 HF axioms =) H ` A" | Ind: "A 2 induction axioms =) H ` A" | MP: "H ` A IMP B =) H’ ` A =) H [ H’ ` B" | Exists: "H ` A IMP B =) atom i ] B =) 8 C 2H. atom i ] C =) H ` (Ex i A) IMP B" Note that the existential rule is subject to the condition that the bound variable, i, is fresh with respect to B and the formulas in H . The definitions The variable i must be of boolean axioms, etc., are taken from ?). He formalised a simpler inference system, with theorems of the H fresh for B and form ` A, but introducing H allows a proof of the deduction theorem and the derivation of a sort of sequent calculus. This is essential if we are to conduct proofs in this formal calculus.
  • 15. Early Steps in the HF Calculus ✤ the deduction theorem (yielding a sequent calculus) ✤ derived rules to support explicit formal proofs ✤ ✤ ✤ for defined connectives, including ∧ ➝ ∀ for equality, set induction, … definitions and proofs for subsets, extensionality, foundation and natural number induction
  • 16. states that the quantified variable (z ) must be fresh for the terms t and u. MemI: and in contrast to some treatments, the bound variable is a In other words,"ss fm (Var i IN Var j)" | DisjI: "ss fm A =) ss than being fm (A OR B)" parameter of the definition ratherfm B =) ssfixed; however, the choice of z | a↵ect "ss fm A =) ss fm B =) ss fm (A AND to cannot ConjI:the denotation of the right-hand side, thanksB)" quotienting. | ExI: "ss fm A =) ss fm (Ex i A)" | All2I: "ss "fm =) atom j ] (i,A) =) ss fm (All2 i (Var j) A)" inductive ss fm ::fm A ) bool" where Σ Formulas MemI: "ss fm (Var i IN Var j)" One advantage A =) ss fm B =) ss fm (A OR B)" | DisjI: "ss fmof formal proof is that these conditions are immediately ev| ConjI: "sstheyAmay not fm B =) from an informal presentation. ?) does ident, when fm =) ss be clear ss fm (A AND B)" | ExI: "ss fm A condition (Ex i A)" not impose the last =) ss fm (on the bound of a universal quantifier), but it statesStrict fm A =) atom j ] containto reach the second j) A)" terms t the | All2I: "ss $ quantified variable(i,A) =) bevariablesiand areand u.basis fm (All2 greatlythat the formulas only(z )needed ssfresh for the (Var incompletesimplifies the main induction must for In other words, and are only interested treatments, the boundincompleteness ness theorem. (If we in contrast to somein second incompleteness theorem. formalising the first variable is a the the definition is that the parameter of of formal proofrather than being fixed; are immediately evOne advantagemain induction of these conditionshowever, the choice of z theorem, we can use a more generous notion of ⌃ formula, allowing atomic ident, when they may notstill derive the general caseto quotienting. be clear from an informal thanks of ?) does cannot a↵ect the denotation of the right-hand side,presentation. $ formulas. formulas andWe can their negations over arbitrary terms.) Formally, a ⌃ formula is not impose the last condition (on the bound of a universal quantifier), but it defined to be any formula) bool" where inductive ss the main that can be proved equivalent (in the HF calculus) greatly simplifiesfm :: "fminduction needed to reach the second incompleteto a strict ⌃"ss fm (Var interested in MemI: formula: ness theorem. (If we are onlyi IN Var j)"formalising the first incompleteness | DisjI: OR B)" theorem, we fm "ss fma Amore generous =) ss of ⌃✓formula, allowing atomic B)" can use! ( =) ss fm BB notion fm (A supp A & {} ` A IFF "Sigma A 9 B. ss fm & supp B | ConjI:their negations over arbitrary ss fm (AFormally, a ⌃ formula is "ss fm A =) ss fm B =) terms.) AND B)" formulas and | to be "ss fm The condition suppAB that can A(Ex i A)" ✓ ss fm proved equivalent that every variable defined ExI: any formula=) supp be essentially means (in the HF calculus) free to in strict ⌃ also be free =)A . After a certain =) ss fm (All2 iit(Var j) A)"to a | All2I: "ss fm A in atom j ] (i,A) amount of e↵ort, is possible B must formula: derive the expected properties of ⌃ formulas and ultimately to reach a key One advantage of formal proof & supp B ✓ supp A & are immediately "Sigma fm A ! ( 9 B. ss fm B is that these conditions {} ` A IFF B)" evresult based they may not be clear from an informal presentation. ?) does on this concept: ident, when The condition the last conditionA(on the bound of a universal quantifier),free it not impose supp B ✓ supp essentially means that every variable theorem " [[Sigma fm A; ground fm A; eval fm e0 A ]] =) {} ` A" but in B must also be free inmain induction needed to reach the second incompletegreatly simplifies the A . After a certain amount of e↵ort, it is possible to derive the expected we are only of ⌃ formulas and ultimately to incompleteness ` If A is a true ⌃ sentence, then `A. This result reduces the task of a key ness theorem. (If properties interested in formalising the first reach proving result basedformal calculus to proving that A holds formula, allowing atomic in A in the we can use a more generous notion of ⌃ (written eval fm e0 A ) theorem, on this concept: True $ formulas are theorems! formulas and their negations over fm A; Isabelle/HOL’s imp thm: " [[Sigmaarbitrary terms.) A; eval a e0 A ]] =) theorem Sigma fm native higher-order logic.ground fm Formally, fm ⌃ formula is
  • 17. Coding Terms and Formulas ✤ must first translate from nominal to de Bruijn format ✤ ✤ ✤ the actual coding is a simple recursive map: ⌜0⌝=0 , ⌜xk⌝=k, ⌜x ◁ y⌝= 〈⌜◁⌝, ⌜x⌝, ⌜y⌝〉, … also define (in HF) predicates to recognise codes ✤ abstraction over a variable — needed to define Form(x), the predicate for formulas ✤ substitution for a variable
  • 18. details. Many other authors prefer to simplify matters via repeated appeals where "MakeForm y u w ⌘ to Church’s thesis. Even the detailed presentations mentioned above omit any y = q Disj u w _ y = q Neg u _ demonstration that the definitions are correct. The proof formalisation condi( 9 v u’. AbstForm v 0 u u’ ^ y = q Ex u’)" tion for the provability predicate (written PfP below) is typically stated with a minimum of code of a formula constructed from existing formulas u and v Thus y is thejustification: provided y codes the disjunction u _ v, the negation ¬u or the existential formula theorem proved iff proved Pf: "{} ` ↵ ! {} ` PfP p↵q" 9(u’), where u’ has been obtained by abstracting u over some variable, v. The predicate interplay of these various points can be seen below: The AbstForm performs de Bruijn abstraction over a formula; its definition is complicated, and omitted here. Note that the codes of quantified formulas do definition the names not mention MakeForm ::of"hf ) hf ) hf ) bool" bound variables. where "MakeForm is given ⌘ a higher-order logic formula, and therefore at the This predicate y u w by y = q Disj u w _ y = q Neg u _ level of the meta-theory. Working at this level eliminates the need to construct ( 9 v u’. AbstForm v 0 u u’ ^ y = q Ex u’)" HF proofs, and most of the correctness properties we need can be proved in this manner. However, in formula constructed from existing formulas u and v Thus y is the code of a % w, or y = ¬ u, ordiagonalisation argument and y = u order to perform the y = (∃v) u exhibit theyundecidable formula, we need a negationof u or the existential formula provided codes the disjunction u _ v, the version ¬ every coding predicate as with an explicit abstraction over on both levels: an HF formula. Therefore, obtained by abstracting definedon u variable, v. The 9(u’), where u’ has been each predicate must be ustep some predicate AbstForm performs de Bruijn abstraction over a formula; its definition nominal primrec MakeFormP :: "tm ) tm ) tm ) fm" is complicated, and omitted here. Note that the codes of quantified formulas do where " [[atom v ] (y,u,w,au); atom au ] (y,u,w) ]] =) not mentionythe wnames of bound variables. MakeFormP u = This predicate is given by higher-order logic formula, and therefore at the y EQ Q Disj u w OR y EQ QaNeg u OR level of the meta-theory. Workingv) Zero level eliminates theEQ Q Ex (Var au)))" Ex v (Ex au (AbstFormP (Var at this u (Var au) AND y need to construct HF proofs, and most of the correctness properties we need can be proved in this manner. However, in order to perform the diagonalisation argument and The “official” version as an HF formula, every boolean exhibit the undecidable formula, we need a version of not acoding predicate as an HF formula. Therefore, each predicate must be defined on both levels: 3 Example: Making a Formula nominal primrec MakeFormP :: "tm ) tm ) tm ) fm"
  • 20. … And Proof Predicates Sent Equality_ax Axiom Prf HF_ax ModPon Special_ax Exists Induction_ax Subst (a sequence of proof steps, and finally...) Pf
  • 21. quantified variables. Here there are only two, but to define AbstForm requires As bound variables. the definition of Subset, constraints are and messy, and 12 we saw above in The necessary declarations are lengthy required on all quantified variables. on thethere are only two, but torun very slowly), requires Here put a heavy burdenabove innominal package (proofs , define AbstForm but the on all As we saw The necessary declarations are constraints messy, and the definition of 12 bound variables. to rename explicit bound Subset lengthy and are required alternative of having variables unattractive. quantified variables.nominal package only two, but alsoslowly), but the requires Here there are (proofs run is to define AbstForm put weheavy above inon the burden the definition of Subset, constraints are required on all As a sawbound variables. The necessary declarations very lengthy and messy, and 12 of having alternative variables. to rename explicit bound variables isare AbstForm requires also quantified a heavy burden on the nominal package (proofs run very slowly), but the Here there are only two, but to define unattractive. put 3 bound variables. having to rename explicit bound lengthy andalso unattractive. proof development: first 12 The Isabelle/HOL declarations are variables is messy, and alternative of The necessary put incompleteness theorem. development: first heavy Isabelle/HOL proof 3 aThe burden on the nominal package (proofs run very slowly), but the alternative of having to rename explicit bound variables is also unattractive. incompleteness K(p q) = p (pproof development: first 3 The Isabelle/HOL q)q ` theorem. Steps to the First Theorem incompleteness theorem. first We lemmaaprove KRP: "{}such that i qK(pq q) = stated, but it is neither a senneedincompleteness pVar ` pA easily p (p q) q" function K ` KRP theorem. The property of being single-valued is pA(i::= pA q)q ` pVar p pA(i::= 3 The KRP: "{} ` KRP K(p i q=pA q (p development: lemma prove Isabelle/HOLq)proof q)q pA q) q" ✤ ✤ ✤ ✤ tence nor a ⌃ formula. Proving`thisq) = p (pq q)qq pA(i::=600 lines of explicit lemmano function symbols. Instead, defineit q) q" prove KRP:single-valuedpVar requires about pA is neither a KRP: "{}K(p result i pA KRP is easily stated, but a relation, sen… but The propertyin the HF calculus, verifying that substitution over terms or we have of being ` reasoning steps tence nor a ⌃ formula. Proving this result requires about 600 lines of explicit formulasprove KRP: "{} result. single-valued is easily stated, lemma yields a unique The in the HF calculus, verifying that pA q) q" being reasoning steps property`ofKRP pVar i q pA q pA(i::=substitution but it terms or a senover is neither tence nor a Proving requires formulasKRP unique: ⌃ formula. y, KRP v this resulty’ proof about 600 lines of explicit yields a lemma property unique result. "{KRP single-valued isx easily` v x y’} HF EQ but it is neither a senProving its functional behaviour takes 600 stated, y" steps. The reasoning of being the HF calculus, verifying that substitution over terms or steps in tence nor a unique: "{KRP now reached by y’} ` y’ about 600 lines of explicit lemma formulas lemmaaProving result.v xthe standard argument. The obtains v x y, KRP The KRP ⌃ formula. isunique this result requires EQ y" diagonal yields reasoning steps in form of calculus, quantification, and is equivalent terms or syntax represents a the HF existential verifying that substitution over to 9 .... Finally, the yields alemma isresult. v x by the standard argument.y" obtains The diagonalKRP unique: "{KRP lemma unique now formulas diagonal lemma:reached y, KRP v x y’} ` y’ EQ The syntax represents a form of existential quantification, and is equivalent to 9 .... lemma diagonal: lemma KRPThe diagonal lemma y, now reached by the standard argument. The obtains x is KRP v x y’} "supp obtains unique: "{} ` v IFF ↵(i::= p q)" ` y’ EQ y" supp ↵ - {atom i}" where "{KRP = lemma syntax represents a form of existential quantification, and is equivalent to 9 .... diagonal: obtains part lemma is now reached by q)" "supp supp - i}, states TheThe diagonal of the conclusion, namelythe standard argument.↵The obtains second where "{} ` IFF ↵(i::= p supp = supp= ↵ - {atom {atom i}" lemma diagonal: syntax represents a form the formula are those of ↵ and the exception 9 .... that second part of theof of existential quantification, supp is↵equivalenti},of i ; it the free variables conclusion, namely supp = with - {atom to states The obtains where "{} ` IFF ↵(i::= p q)" "supp = supp ↵ - {atom i}" is necessary in order to show that the undecidable formula is actually a sentence.
  • 22. theorem Goedel I: assumes Con: " ¬ {} ` Fls" obtains where "{} ` IFF Neg (PfP p q)" " ¬ {} ` " " ¬ {} ` Neg " "eval fm e " "ground fm " proof obtain where "{} ` IFF Neg ((PfP (Var i))(i::= p q))" and [simp]: "supp = supp (Neg (PfP (Var i))) - {atom i}" by (metis SyntaxN.Neg diagonal) hence diag: "{} ` IFF Neg (PfP p q)" by simp hence np: " ¬ {} ` " by (metis Con Iff MP same Neg D proved iff proved Pf) hence npn: " ¬ {} ` Neg " using diag by (metis Iff MP same NegNeg D Neg cong proved iff proved Pf) moreover have "eval fm e " using hfthm sound [ where e=e, OF diag] by simp (metis Pf quot imp is proved np) moreover have "ground fm " by (auto simp: ground fm aux def) ultimately show ?thesis by (metis diag np npn that) qed
  • 23. Steps to the Second Theorem ✤ Coding must be generalised to allow variables in codes. ✤ ✤ ✤ ✤ ⌜x ◁ y⌝ = 〈⌜◁⌝, ⌜x⌝, ⌜y⌝〉 ⎣x ◁ y⎦V = 〈⌜◁⌝, x, y〉 codes of variables are integers Variables must be renamed, with the intent of creating “pseudoterms” like 〈⌜◁⌝, Q x, Q y〉. Q is a magic function: Q x = ⌜t⌝ where t is some canonical term denoting the set x.
  • 24. Complications ✤ Q must be a relation. ✤ Function symbols cannot be added… ✤ Sets do not have an easily defined canonical ordering. ✤ QR(0,0) ✤ QR(x,x’), QR(y,y’) ⟹ QR(x ◁ y, 〈⌜◁⌝, x’, y’〉)
  • 25. ?) eliminates the ambiguity implicit in (5) above by appealing to a total ordering, <, on the HF universe. The di culty is how to define this ordering ´ within the HF calculus. Swierczkowski develops the theory, including a definition by recursion on the rank of a set, but it does not look easy to formalise in HF. P Another approach is to define the function f : HA ! N such that f (x) = {2f (y) | y 2 x}. Then we can define x < y () f (x) < f (y). Again, the e↵ort to formalise this theory in HF may be simpler than that needed to formalise the Chinese remainder theorem, but it is still considerable. One of the Final Lemmas QR(x, x0 ), QR(y, y 0 ) ` x 2 y ! Pf bx0 2 y 0 c{x0 ,y0 } QR(x, x0 ), QR(y, y 0 ) ` x ✓ y ! Pf bx0 ✓ y 0 c{x0 ,y0 } QR(x, x0 ), QR(y, y 0 ) ` x = y ! Pf bx0 = y 0 c{x0 ,y0 } ✤ ✤ 2 ; () ? The first two require zsimultaneous induction, yielding the third. z 2 x C y () z 2 x _ z = y ; symbols % Similar proofs for the ✓ z () > & ∃ and bounded ∀. x C y ✓ z () x ✓ z ^ y 2 z ✤ The proof in the HF x = y () needs^ y ✓ x 450 lines. calculus x ✓ y under The point of all this is that (??) and (??) can be proved by a simultaneous ✤ Fills a major gap in various presentations, including S!wierczkowski’s. induction: QR(x, x0 ), QR(y, y 0 ) ` (x 2 y ! Pf bx0 2 y 0 c{x0 ,y0 } ) ^ (x ✓ y ! Pf bx0 ✓ y 0 c{x0 ,y0 } )
  • 26. =) Pf bx0 ✓ y 0 c{x0 ,y0 } ^ Pf bx0 2 y 0 c{x0 ,y0 } 1 2 1 2 () Pf bx0 C x0 ✓ y 0 c{x0 ,x0 ,y0 } 1 2 1 2 theorem Goedel II: assumes Con: " ¬ {} ` Fls" shows " ¬ {} ` Neg (PfP pFls q)" proof from Con Goedel I obtain where diag: "{} ` IFF Neg (PfP p q)" " ¬ {} ` " and gnd: "ground fm " by metis have "{PfP p q} ` PfP pPfP p qq" by (auto simp: Provability ground fm aux def supp conv fresh) moreover have "{PfP p q} ` PfP pNeg (PfP p q) q" apply (rule MonPon PfP implies PfP [OF gnd]) apply (auto simp: ground fm aux def supp conv fresh) using diag by (metis Assume ContraProve Iff MP left Iff MP left’ Neg Neg iff) moreover have "ground fm (PfP p q)" by (auto simp: ground fm aux def supp conv fresh) ultimately have "{PfP p q} ` PfP pFls q" using PfP quot contra by (metis (no types) anti deduction cut2) thus " ¬ {} ` Neg (PfP pFls q)" by (metis Iff MP2 same Neg mono cut1 diag) qed
  • 27. What Did We Learn? ✤ Some highly compressed proofs were finally made explicit. ✤ The entire proof development can be examined interactively. ✤ The nominal package can cope with very large developments… (BUT: performance issues, some repetitive notation, complications in accepting function definitions) ✤ <9 months for the first theorem, a further 4 for the second ✤ Under 16 000 lines of proof script in all.
  • 28. Conclusions ✤ the first-ever machine formalisation of Gödel’s second incompleteness theorem ✤ using both nominal and de Bruijn syntax for bound variables ✤ within an axiom system for hereditarily finite set theory ✤ conducted using Isabelle/HOL. Many thanks to Christian Urban for help at critical points! Also Jesse Alama, Dana Scott.