Tuesday 27 October 2009

Cantors Diagonalization Proof in Coq


Axiom b : (nat -> nat) -> nat.
Axiom b' : nat -> (nat -> nat).

Axiom bijection : forall f, b' (b f) = f.
Axiom bijection' : forall x, b (b' x) = x.

Definition Stream A := nat -> A.
Definition Map {A B} (f : A -> B) : Stream A -> Stream B :=
fun sa =>
fun i =>
f (sa i).

Definition l : Stream (nat -> nat) -> Stream nat :=
Map b.
Definition l' : Stream nat -> Stream (nat -> nat) :=
Map b'.

Definition isnt x :=
match x with
| O => S O
| S _ => O
end.

Theorem isn't_isnt {x} : isnt x <> x.
destruct x; discriminate.
Qed.

Definition Diagonal {A} (s : Stream (Stream A)) : Stream A:=
fun i => s i i.

Definition Diagonalization (sequences : Stream (nat -> nat)) : Stream nat :=
Map isnt (Diagonal sequences).

Theorem isn't_in {sequences} : forall i, Diagonalization sequences <> sequences i.
intros; intro.
pose (f_equal (fun f => f i) H) as H'.
unfold Diagonalization in H'.
unfold Diagonal in H'.
unfold Map in H'.
revert H'.
apply isn't_isnt.
Qed.

Definition pink_elephant : Stream nat := fun i => i.

Definition sequence_enumeration : Stream (nat -> nat) := l' pink_elephant.

Theorem every_sequence : forall s, exists i, s = sequence_enumeration i.
intro s; exists (b s).
unfold sequence_enumeration.
unfold pink_elephant.
unfold l'.
unfold Map.
rewrite bijection.
reflexivity.
Qed.

Theorem Cantor : False.
destruct (every_sequence (Diagonalization sequence_enumeration)) as [i iPrf].
pose (@isn't_in sequence_enumeration i).
contradiction.
Qed.



Saturday 24 October 2009

Short Note on Semantics

With all curious derivations that come from the forests of To Mock a Mockingbird, and in the intrests of deepening understanding of type theory (in particular CIC) we considered axiomatizing the SK calculus in Coq.



Module Type SK.
(* axiomatization of the theory *)
Parameter SK : Set.

Parameter app : SK -> SK -> SK.
Infix "*" := app.

Parameter S : SK.
Parameter K : SK.

Axiom S_prop : forall a b c, S * a * b * c = (a * c) * (b * c).
Axiom K_prop : forall a b, K * a * b = a.
End SK.


It is not defined as an inductive type (as you might usually do) because there is no normal form for SK that we can compute terms into inside Coq (we would have to equate diverging terms with an "omega" symbol for a start, which is well known to be a partial operation). It would be possible to use a non-normal forma and quotient it out with a special (undecidable) equivalence relation, but we do not this.. Instead the axioms of the theory are given, One immediately notices that there is a trivial model of these axioms:



Module SK_trivial_model : SK.
(* one model (giving non-contradiction of the theory) *)
Definition SK := unit.

Definition app := fun (_ _ : unit) => tt.
Infix "*" := app.

Definition S := tt.
Definition K := tt.

Theorem S_prop : forall a b c, S * a * b * c = (a * c) * (b * c).
reflexivity.
Qed.

Theorem K_prop : forall a b, K * a * b = a.
destruct a; reflexivity.
Qed.
End SK_trivial_model.


(This implementation mostly serves to show the syntax and so on of how to construct a model of a theory inside Coq).

Anyway, we eliminate trivial models by saying that S <> K.



  Axiom SK_prop : S <> K.


We could now develop the theory of combinator calculus, prove things like S <> I and so forth.



Module SK_theory <: SK.
Include Type SK.

(* Development of the theory of SK calculus, valid for all models *)

Ltac sk := repeat (rewrite S_prop || rewrite K_prop).

Definition I := S * K * S.
Theorem I_prop : forall x, I * x = x.
intros; unfold I; sk; reflexivity.
Qed.
Ltac ski := repeat (sk || rewrite I_prop).

Theorem IK : I <> K.
Definition F x y b := b * K * I * x * y.

intro Eq; generalize (f_equal (F K S) Eq); unfold F; ski.
apply SK_prop.
Qed.

Definition U := S * I * I.
Theorem U_prop : forall x, U * x = x * x.
intros; unfold U; ski; reflexivity.
Qed.

Ltac skiu := repeat (ski; rewrite U_prop).
End SK_theory.


As fun as that would be, I don't believe you could ever reach a contradiction by assuming the SK theory. Yet no model exists inside Coq!

This is a rather confusing and troubling state, but Russell O'Connor mentioned to me the possibility of giving an interpretation of CIC into a theory that has a Halting-Oracle and that solves it! This is surely a valid interpretation of CIC, it's just got this SK theory added in with the usual stuff - the halting oracle would be used to normalize SK terms and prove satisfy our theory. So despite the non-existence of a model of SK inside Coq: It appears we can reason about the SK calculus directly. (Rather than a halting oracle you might have considered a special equivalence relation being interpreted for the congruence of SK objects but lets pretend Coq has Axiom K so that need not be considered).

This is all very interesting to me because it seems to suggest mathematics existing outside of the formal theory we work in or study. Something I never believed in - It's not completely clear if this is profound or if I am just going soft.