An MIU Decision Procedure in Lean #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
The MIU formal system was introduced by Douglas
Hofstadter in the first chapter of his 1979 book,
Gödel, Escher, Bach.
The system is defined by four rules of inference, one axiom, and an alphabet of three symbols:
M
, I
, and U
.
Hofstadter's central question is: can the string "MU"
be derived?
It transpires that there is a simple decision procedure for this system. A string is derivable if
and only if it starts with M
, contains no other M
s, and the number of I
s in the string is
congruent to 1 or 2 modulo 3.
The principal aim of this project is to give a Lean proof that the derivability of a string is a decidable predicate.
The MIU System #
In Hofstadter's description, an atom is any one of M
, I
or U
. A string is a finite
sequence of zero or more symbols. To simplify notation, we write a sequence [I,U,U,M]
,
for example, as IUUM
.
The four rules of inference are:
- xI → xIU,
- Mx → Mxx,
- xIIIy → xUy,
- xUUy → xy,
where the notation α → β is to be interpreted as 'if α is derivable, then β is derivable'.
Additionally, he has an axiom:
MI
is derivable.
In Lean, it is natural to treat the rules of inference and the axiom on an equal footing via an
inductive data type derivable
designed so that derivable x
represents the notion that the string
x
can be derived from the axiom by the rules of inference. The axiom is represented as a
nonrecursive constructor for derivable
. This mirrors the translation of Peano's axiom '0 is a
natural number' into the nonrecursive constructor zero
of the inductive type nat
.
References #
- Jeremy Avigad, Leonardo de Moura and Soonho Kong, Theorem Proving in Lean
- Douglas R Hofstadter, Gödel, Escher, Bach
Tags #
miu, derivable strings
Declarations and instance derivations for miu_atom
and miustr
#
- M : miu.miu_atom
- I : miu.miu_atom
- U : miu.miu_atom
The atoms of MIU can be represented as an enumerated type in Lean.
Instances for miu.miu_atom
- miu.miu_atom.has_sizeof_inst
- miu.miu_atom_inhabited
- miu.miu_atom.has_repr
- miu.miustr.has_mem
The annotation @[derive decidable_eq]
above assigns the attribute derive
to miu_atom
, through
which Lean automatically derives that miu_atom
is an instance of decidable_eq
. The use of
derive
is crucial in this project and will lead to the automatic derivation of decidability.
We show that the type miu_atom
is inhabited, giving M
(for no particular reason) as the default
element.
Equations
miu_atom.repr
is the 'natural' function from miu_atom
to string
.
Equations
- miu.miu_atom.U.repr = "U"
- miu.miu_atom.I.repr = "I"
- miu.miu_atom.M.repr = "M"
Using miu_atom.repr
, we prove that ``miu_atomis an instance of
has_repr`.
Equations
- miu.miu_atom.has_repr = {repr := λ (u : miu.miu_atom), u.repr}
For simplicity, an miustr
is just a list of elements of type miu_atom
.
Equations
Instances for miu.miustr
For display purposes, an miustr
can be represented as a string
.
Equations
- miu.miustr.mrepr (c :: cs) = c.repr ++ miu.miustr.mrepr cs
- miu.miustr.mrepr list.nil = ""
Equations
- miu.miurepr = {repr := λ (u : miu.miustr), u.mrepr}
In the other direction, we set up a coercion from string
to miustr
.
Equations
- miu.lchar_to_miustr (c :: cs) = let ms : miu.miustr := miu.lchar_to_miustr cs in miu.lchar_to_miustr._match_1 ms c
- miu.lchar_to_miustr list.nil = list.nil
- miu.lchar_to_miustr._match_1 ms _x = list.nil
- miu.lchar_to_miustr._match_1 ms 'U' = miu.miu_atom.U :: ms
- miu.lchar_to_miustr._match_1 ms 'I' = miu.miu_atom.I :: ms
- miu.lchar_to_miustr._match_1 ms 'M' = miu.miu_atom.M :: ms
Equations
- miu.string_coe_miustr = {coe := λ (st : string), miu.lchar_to_miustr st.data}
Derivability #
- mk : miu.derivable ↑"MI"
- r1 : ∀ {x : miu.miustr}, miu.derivable (x ++ [miu.miu_atom.I]) → miu.derivable (x ++ [miu.miu_atom.I, miu.miu_atom.U])
- r2 : ∀ {x : list miu.miu_atom}, miu.derivable (miu.miu_atom.M :: x) → miu.derivable (miu.miu_atom.M :: x ++ x)
- r3 : ∀ {x y : miu.miustr}, miu.derivable (x ++ [miu.miu_atom.I, miu.miu_atom.I, miu.miu_atom.I] ++ y) → miu.derivable (x ++ miu.miu_atom.U :: y)
- r4 : ∀ {x y : miu.miustr}, miu.derivable (x ++ [miu.miu_atom.U, miu.miu_atom.U] ++ y) → miu.derivable (x ++ y)
The inductive type derivable
has five constructors. The nonrecursive constructor mk
corresponds
to Hofstadter's axiom that "MI"
is derivable. Each of the constructors r1
, r2
, r3
, r4
corresponds to the one of Hofstadter's rules of inference.