Course5: Proofs as programs

First part of the lecture:

back to dependent types for some complements!


Using Coq as a proof assistant.

At last, we shall start using Coq as a proof assistant. But before really doing so, we shall see that in fact, we already have all we need to do this with what you learnt from the start of the semester!
The present section shall thus be seen as a pedagogical introduction to modelling from scratch proofs as Coq programs, that is really putting Curry-Howard correspondence at work, before really considering how this is actually done in Coq.
In order to prevent mixing Coq actual features with our own development, I use french terms in this introductory part, which will also be an occasion for practicing french for some of you!
To start with, we set implicit arguments to lighten notations and we define Formulas as an alias for Set. (this will be updated later!)

Set Implicit Arguments.

Definition Formula := Set.

In the following, we shall introduce some types and type constructions in Formula as well as some Coq terms which shall model inference rules on them.
We shall follow the outline:
  • true
  • false
  • implication
  • negation
  • conjunction
  • disjunction
  • excluded-middle
  • quantifiers
  • equality
  • first proofs

Let us consider the following inductive types:
with its single constructor (introduction rule):

Inductive VRAI : Formula :=
| VRAIintro.

Check VRAIintro.
VRAIintro : VRAI
VRAI has essentially one proof (inhabitant): its single constructor.

with no constructor (no introduction rule) and a generic destructor, eliminating to any formula:

Inductive FAUX : Formula :=
.

Having no constructor, FAUX has no introduction rule.

Definition FAUXelim (A: Formula) (f: FAUX) : A :=
match f with
end.

Notice the odd shape of this pattern-matching with no clause for matching... but indeed FAUX has no constructor!

Check FAUXelim.

FAUXelim : A : Formula, FAUX A
FAUXelim corresponds to the elimination of the absurd, also known as ex falso sequitur quodlibet.

with its constructor / introduction rule (if from assuming A, I can deduce B, then I can deduce A => B), and its destructor / elimination rule (Modus Ponens):

Inductive IMP (A B : Formula) : Formula :=
|IMPintro: (AB) → IMP A B.


Check IMPintro.

IMPintro : A B : Formula, (A B) IMP A B

Definition IMPelim (A B: Formula) (i: IMP A B) (a: A) : B :=
match i with
|IMPintro ff a
end.

Check IMPelim.

IMPelim : A B : Formula, IMP A B A B
Notice that an alternative, would have been to define IMP simply as an alias for :

Definition IMP' (A B: Formula) : Formula := AB.

We will come back to this later today.

with its introduction rule and its elimination rule (from not A and A, deduce false):

Inductive NEG (A : Formula) : Formula :=
|NEGintro : (AFAUX) → NEG A.


Check NEGintro.
NEGintro : A : Formula, (A FAUX) NEG A

Definition NEGelim (A : Formula) (n: NEG A) : AFAUX :=
match n with
|NEGintro ff
end.

Check NEGelim.
NEGelim : A : Formula, NEG A A FAUX
Negation could also be considered as a defined connective:

Definition NEG' (A : Formula) : Formula := IMP A FAUX.

with its constructor (introduction rule) and its two destructors (elimination rules) :

Inductive ET (A B : Formula) : Formula :=
| ETintro : A → (BET A B).

Check ETintro.
ETintro : A B : Formula, A B ET A B

Definition ETelim1 (A B: Formula) (c: ET A B) : A :=
match c with
|ETintro a ba
end.

Definition ETelim2 (A B: Formula) (c: ET A B) : B :=
match c with
|ETintro a bb
end.

Check ETelim2.
ETelim2 : A B : Formula, ET A B B

Now that we have disjunction, equivalence can be obtained as a defined connective:

Definition EQUIV (A B: Formula) := ET (IMP A B) (IMP B A).

with its two constructors (introduction rule) and its elimination rule:

Inductive OU (A B : Formula) : Formula :=
|OUintro1 : AOU A B
|OUintro2 : BOU A B.

Arguments OUintro1 {A B}.
Arguments OUintro2 {A B}.

Check @OUintro1.
OUintro1 : A B : Formula, A OU A B

Definition OUelim {A B C: Formula}
(d: OU A B) (e: AC) (f: BC) :=
match d with
|OUintro1 ae a
|OUintro2 bf b
end.

Check @OUelim.
OUelim : A B C : Formula, OU A B (A C) (B C) C

Shall a proof of a disjunction only be obtained from a proof of the left disjunct or from the right disjunct? Is there no third option?
The excluded-middle is a third option, if you want to accept it.
But there is no computational counter-poart to it: one shall add it as an axiom, that is as an additional constant we declare.
Do you remember how to do it?


Section ExcludedMiddle.
Variables A B: Formula.
Variable TiersExclus : (OU A (NEG A)).

Definition PeirceLaw : IMP (IMP (IMP A B) A) A :=
IMPintro (fun (f : IMP (IMP A B) A) ⇒ match TiersExclus with
| OUintro1 PAPA
| OUintro2 (NEGintro fnegA) ⇒ match f with
|IMPintro f'f' (IMPintro (fun (a: A) ⇒
match (fnegA a) return B with
end)) end
end).

What is this term PeirceLaw ??

End ExcludedMiddle.

Check PeirceLaw.

We now introduce formula constructions and deduction schemas for quantifiers. Both for the universal and existential quantifier, we provide two versions of the quantifiers:
  • an arithmetical quantification and
  • a general quantification over any type in Set.
Forall qantifier

Inductive POURTOUTN (B: natFormula) : Formula :=
|PTNintro : ( (x : nat), B x) → POURTOUTN B.

Arguments PTNintro {B}.

Check @PTNintro.
PTNintro : B : nat Formula, ( x : nat, B x) POURTOUTN B

Definition POURTOUTNelim {B: natFormula}
(f: POURTOUTN B) (t: nat) : B t :=
match f with
|PTNintro gg t
end.

Check @POURTOUTNelim.
POURTOUTNelim : B : nat Formula, POURTOUTN B t : nat, B t

Inductive POURTOUT (A : Set) (B: AFormula) : Formula :=
|PTintro : ( (x : A), B x) → POURTOUT B.

Arguments PTintro {A B}.

Definition POURTOUTelim {A: Set} {B: AFormula}
(f: POURTOUT B) (t: A) : B t :=
match f with
|PTintro gg t
end.

Exists quantifier

Inductive ILEXISTEN (B: natFormula) : Formula :=
|IENintro : x: nat, B xILEXISTEN B.

Arguments IENintro {B}.

Check @IENintro.
IENintro : (B : nat Formula) (x : nat), B x ILEXISTEN B
Be careful in reading this type!!
Let us add some parentheses:
IENintro : (B : nat Formula) (x : nat), (B x ILEXISTEN B)

Definition ILEXISTENelim {B: natFormula} {C: Formula}
(e: ILEXISTEN B) (f: x: nat, B xC) : C :=
match e with
|IENintro t pf t p
end.

Check @ILEXISTENelim.
ILEXISTENelim : (B : nat Formula) (C : Formula), ILEXISTEN B ( x : nat, B x C) C
Again, it is useful to add some parentheses to parse the above type correctly:
ILEXISTENelim : (B : nat Formula) (C : Formula), ILEXISTEN B ( x : nat, (B x C)) C

Inductive ILEXISTE (A: Set) (B: AFormula) : Formula :=
|IEintro : x: A, B xILEXISTE B.

Arguments IEintro {A B}.

Definition ILEXISTEelim (A: Set) {B: AFormula} {C: Formula}
(e: ILEXISTE B) (f: x: A, B xC) : C :=
match e with
|IEintro t pf t p
end.


The only introduction rule for equality is the reflexivity of equality. Using equality amounts to substituting equals by equals and this is dealt with through dependent pattern-matching.

Inductive EQN (x : nat) : natFormula :=
| EQNintro : EQN x x.

Check EQNintro.
EQNintro : x : nat, EQN x x

Definition EQNelim (A: natFormula) (m n : nat):
EQN m nA mA n :=
fun (Heq: EQN m n) ⇒ fun (x: A m) ⇒
match Heq in (EQN _ a) return A a with
|EQNintro _x
end.

Check EQNelim.
EQNelim : (A : nat Set) (m n : nat), EQN m n A m A n

Inductive EQ (A: Set) (x : A) : AFormula :=
| EQintro : EQ x x.

Definition EQelim (B: Set) (A: BFormula) (m n : B):
EQ m nA mA n :=
fun (Heq: EQ m n) ⇒ fun (x: A m) ⇒
match Heq in (EQ _ a) return A a with
|EQintro _x
end.

One can state and prove the first properties about equality :

Definition EQNrefl:
POURTOUTN (fun xEQN x x) :=
PTNintro (fun xEQNintro x).

Definition EQNtrans : POURTOUTN (fun x
POURTOUTN (fun y
POURTOUTN (fun z
IMP (EQN x y)
(IMP (EQN y z) (EQN x z))))) :=
PTNintro (fun x
PTNintro (fun y
PTNintro (fun z
IMPintro (fun (H : EQN x y) ⇒
IMPintro (fun (G: EQN y z) ⇒
match G in (EQN _ a) return (EQN x a) with
|EQNintro _H
end))))).

Definition EQNsym:
POURTOUTN (fun x
POURTOUTN (fun y
IMP (EQN x y) (EQN y x))) :=
PTNintro (fun x
PTNintro (fun y
IMPintro (fun (H : EQN x y) ⇒
match H in (EQN _ a) return (EQN a x) with
|EQNintro _EQNintro x
end))).

Check EQNrefl.
Check EQNtrans.
Check EQNsym.

One can state and prove the first properties about equality even over other individual sorts:

Definition EQrefl (A: Set):
POURTOUT (fun (x:A) ⇒ EQ x x) :=
PTintro (fun xEQintro x).

Definition EQtrans (A: Set): POURTOUT (fun (x: A) ⇒
POURTOUT (fun y
POURTOUT (fun z
IMP (EQ x y)
(IMP (EQ y z) (EQ x z))))) :=
PTintro (fun x
PTintro (fun y
PTintro (fun z
IMPintro (fun (H : EQ x y) ⇒
IMPintro (fun (G: EQ y z) ⇒
match G in (EQ _ a) return (EQ x a) with
|EQintro _H
end))))).

Definition EQsym (A: Set):
POURTOUT (fun (x: A) ⇒
POURTOUT (fun y
IMP (EQ x y) (EQ y x))) :=
PTintro (fun x
PTintro (fun y
IMPintro (fun (H : EQ x y) ⇒
match H in (EQ _ a) return (EQ a x) with
|EQintro _EQintro x
end))).

Check EQrefl.
Check EQtrans.
Check EQsym.

Module Notations.
Infix "==" := EQ (at level 70).
Infix "⇒" := IMP (at level 70).
Infix "⇔" := EQUIV (at level 70).

Notation "∀" := POURTOUT (at level 70).
Notation "∀N" := POURTOUTN (at level 70).
Notation "∃" := ILEXISTE (at level 70).
Notation "∃N" := ILEXISTEN (at level 70).

Infix "∧" := ET (at level 50).
Infix "∨" := OU (at level 40).
Notation "¬ L" := (NEG L) (at level 65).
End Notations.
Import Notations.

Check PeirceLaw.

Check EQNtrans.
Check EQrefl.
Check EQtrans.
Check EQsym.

We shall end this introduction with some more complex statements and proof, involving inductive reasoning.
Let us first define two predicates over nat modelling that an natural number is even (resp. odd):

Inductive Even : natFormula :=
| Even0 : Even 0
| EvenSSn : (n: nat), (Even n) → (Even (S (S n))).

Inductive Odd : natFormula :=
| Odd1 : Odd 1
| OddSSn : (n: nat), (Odd n) → (Odd (S (S n))).

It is natural to be willing to prove that any natural is even or odd. Let us do it !

Fixpoint EO (n: nat) : (OU (Even n) (Odd n)) :=
match n with
|0 ⇒ OUintro1 Even0
|1 ⇒ OUintro2 Odd1
|S ((S p) as m) ⇒ match (EO p) with
  |OUintro1 HOUintro1 (EvenSSn H)
  |OUintro2 HOUintro2 (OddSSn H)
end
end.

Check EO.

Definition EvenOrOdd: POURTOUTN (fun (n:nat) ⇒
(OU (Even n) (Odd n))) := PTNintro (EO).

Check EvenOrOdd.
EvenOrOdd : (∀N) (fun n : nat Even n Odd n)

SKIP THIS!!

Definition EO' : (n: nat), (OU (Even n) (Odd n)).
Proof.
fix IHn 1.
intro n.
destruct n.
- apply (OUintro1 Even0).
- destruct n.
  + apply (OUintro2 Odd1).
  + destruct (IHn n) as [He | Ho].
    × apply (OUintro1 (EvenSSn He)).
    × apply (OUintro2 (OddSSn Ho)).
Qed.




Theorem AddComm: POURTOUTN (fun m
POURTOUTN (fun nEQ (m+n) (n+m))).
split. intro m. split. induction m.
- intro n. simpl. induction n.
 + simpl. split.
 + induction n.
  × simpl. split.
  × simpl. simpl in IHn. rewrite IHn. split.
- intro n. simpl. induction n.
 + simpl. simpl in IHm. rewrite IHm. split.
 + simpl. rewrite <- IHn. rewrite IHm. simpl. rewrite IHm. split.
Defined.

Print AddComm.
Check AddComm.
AddComm : (∀N) (fun m : nat (∀N) (fun n : nat m + n == n + m))



Since the start of the semester, we worked with a first universe hierarchy (Type(i), i in |N) (and Set is Type(0).

Check 0.
Check nat.
Check Set.
Alias for the first level Type(0)
Check Type.

The various levels of the hierarchy are used to "type types":

Check nat.
Check natnat.
Inductive Fin : natSet :=
 | Zero n : Fin (S n)
 | Succ n : Fin nFin (S n).
Check (a: nat), Fin aFin a.
Set is Type(0)
Check (A: Set), AA.
Here, Type is Type(1)

Check SetSet.
Here, Type is Type(1)

Check TypeType.
Here, Type Type : Type should be read Type(i) Type(j) : Type(max(i,j)+1 Since Type(i) : Type(i+1) and Type(j) : Type(j+1)

Inductive List (A: Type): Type :=
|Nil: List A
|Cons: AList AList A.
Check List.

Another universe : Prop The universe of logical statement and proofs
For proofs

Check I.
is the canonical proof of True, hence I : True
Check True.
Check Prop.

                Type(...)
                   |
                 Type(3)
                   |
                 Type(2)
                   |
                 Type(1)
               /         \
             /            \
           /               \
Set=Type(0)                Prop
 /   |     \               |   \
nat bool bool->nat       True  False
 |   |       \             |
 O   true    \             I     -  (no close proof of False)
      (fun (x:nat) => 1)           
An upward link between T and S represents the fact that one can derive T : S in Coq type system.
In Coq, the arrow type AB is actually not a primitive construction, it is a particular case of a product x:A,B (Coq syntax x:A, B). When variable x does not occur in B (non-dependent product), we write AB. For a first example of dependent product, see the definition of identity below.
Roughly, the typing rule for a product looks like :
  • If Γ A:Type and Γ+x:A B:Type then Γ ( x:A,B) : Type
In reality, one must take care of the indices of the Type universes : the rightmost index is the max of the other two indices:
  • If Γ A: Type(i) and Γ+x:A B:Type(j) then
Γ ( x:A,B) : Type(max(i,j))
Or, written as an inference rule:
Γ ⊢ A:Type(i)         Γ+x:A ⊢ B:Type(j)
________________________________________(Prod-Type)
  Γ ⊢ (forall x:A,B) : Type(max(i,j))
Prop has some specificities that we shall investigate below in more details.
  • In Set:
    • A specification is any type S of the sort Set
    • A program is any term t the type of which is a specification.
  • In Prop:
    • A proposition or statement is any type P of the sort Prop
    • A proof is any term t the type of which is a proposition.
To sum up: A program is an inhabitant of a specification while a proof is an inhabitant of a proposition.
Dependent product formation rules
Sort triples for dependent product formation rules
The above lines respectively correspond to:
  • simple types
  • impredicativity of Prop
  • dependency
  • higher-order


The user indicates she wants to start a proof with a command such as :
Lemma and_commut : A B : Prop, A B B A.
This command give a name (here and_commut) to the lemma, this will allow later to refer to this lemma later. Instead of Lemma, one can use similar keywords such as Theorem or Proposition or Fact: they are all equivalent form the point of view of Coq and are used to help the human reader to follow a formal proof development, as in common mathematical texts.
After that, the command Proof. is to be used to mark the actual beginning of the proof itself (actually not mandatory but strongly recommanded).

When a proof is started, the proof mode of Coq will display in the higher right part one or several subgoals that are to be proved separately, one by one. These subgoals are essentially sequents of the natural deduction logic, written vertically : the (named) hypothesis comes first on top and the conclusion is below the bar. In the upper part one can also find the variable declarations.
The proof is to be done by tactics, that are orders given by the user to progress in the proof of the current goal, by transforming this goal in various ways. Tactics have lowercase names. For instance, the intro tactic performed on a goal of the form A B will introduce an hypothesis H : A in the context, and replace the conclusion by B. Each inference rule of the natural deduction logic will correspond to one tactic (or more), but some tactics allow to perform more advanced proof steps, for instance solving linear inequalities in the Presburger arithmetic (tactic lia).
Each tactic may generate several subgoals (corresponding to premises of the underlying logical rule), or on the contrary it may completely solve the current goal and hence remove it. The proof is finished when all subgoals are proved. In this case, the command Qed is to be used (from Quod erat demonstrandum) to conclude the proof and leave the proof mode. Here comes a possible complete proof for the previous statement:

Lemma and_commut :
   A B : Prop, ABBA.
Proof.
 intros. split.
 - intros. destruct H. split. assumption. assumption.
 - intros. destruct H. split; assumption.
Qed.

When tactics are separated by dots, Coq will execute them steps-by-steps. One can also use semi-colon ; to chain tactics : the second tactic is applied to all subgoals resulting from the first tactic on the current goal, and so forth. For instance, split;assumption applies assumption on the two subgoals created by split.
Before a tactic, one may optionally write a bullet, i.e. one of the character - or + or ×. These bullets help organizing the proof in a hierarchical way, with delimitation of each sub-part (enforced by Coq).
Such a proof script is to be saved in a file with a .v extension, as done in the previous lectures for instance myproofs.v. Then it can be compiled via the unix command coqc myproofs.v. If the content of this file is correct, then a binary file myproofs.vo is produced, allowing later a fast reload of our proofs (via the Require Import command).


We shall now redo what we did in the first part of this lecture, but using Coq natively defined types. You should experience some déjà vu but there will be some difference (mainly for implication and quantification).

Lemma my_first_proof : True.
Proof.
exact I.
Qed.

Print my_first_proof.

Just about the same as this: Definition my_first_proof_bis : True := I.
Operators, for building more advanced statements

Check natnat.

Check TrueTrue.

Lemma proof2 : TrueTrue.
Proof.
exact (fun _I).
Qed.

or step by step :
intro : the tactic of introduction of -> :
  A |- B
-----------(->intro)
 |- A-> B 

Lemma proof3 : TrueTrue.
Proof.
intro.
Show Proof.
to inspect the proof term being built
assumption.
the axiom rule of natural deduction
Qed.

Print proof3.

Same introduction tactic : intro.

Lemma identity : (A:Prop), AA.
Proof.
intros A a.
assumption.
Qed.

Print identity.

forall and its non-dependent version -> are the only primitive operators.
To introduce them, one has the following introduction tactics: intro / intros / intros ... It is possible to name the hypothesis or quantified variables, otherwise, Coq name them automatically: it can be convenient but may result in less maintainable formal proofs.
On the other hand, in order to used an implicative of universally quantified statement, one uses their elimination: apply H.

Lemma test : (A B : Prop), (AB)->AB.
Proof.
intros A B f a.
apply f.
assumption. apply a.
exact a.
Qed.

Print test.

By default, apply work on the goal to be prove, it is backward-chaining, proving from the conclusion to the hypothesis.
But one can also specify that apply should be used on an hypothesis, reasoning in forward-chaining mode, from the hypotheses to the conclusion:

Lemma test' : (A B : Prop), (AB)->AB.
Proof.
intros A B f a.
apply f in a.
assumption. apply a.
exact a.
Qed.

Print test.
Print test'.

Lemma test'' : (A B : Prop), (AB)->AB.
Proof.
auto.
Qed.

Print test''.

Other operators

Check True.
Print True.
Check I.


Check False.
Print False.
no closed construction

Lemma efql : False (A:Prop), A.
Proof.
intro fa.
intro A.
destruct fa.
elimination of a False hypothesis
Qed.

Print efql.

Compare with the predifined :

Check False_ind.
Print False_ind.

Check False_rect.
Print False_rect.


negation is a shortcut for ...->False

Check ¬True.
Check TrueFalse.

Lemma attempt : ~True. Proof.
unfold "~" in *. intro.
  • introduction is done via the split tactic
  • elimination via destruct ...

Parameter A B : Prop.

Lemma conj : ABBA.
Proof.
 intro H.
 destruct H.
or destruct H as [HA HB].
 split.
 - assumption.
 - assumption.
Qed.

Check and_ind.
Available bullets for structuring a proof script : - + ×
  • introduction is done via the left and right tactics
  • elimination via destruct ...

Lemma disj : ABBA.
Proof.
 intro H.
 destruct H.
or destruct H as [HA | HB].
 - right. assumption.
 - left. assumption.
Qed.

A<->B is just a shortcut for (A->B)/\(B->A)

Lemma disj_equiv: ABBA.
Proof.
split; intro H;
 destruct H; [right | left | right | left]; assumption.
Qed.

Lemma disj_equiv_fail: ABBA.
Proof.
split; intro H;
 destruct H; (right || left); assumption.
Qed.

Lemma disj_equiv': ABBA.
Proof.
split; intro H;
 destruct H; ((right; assumption) || (left; assumption)).
Qed.

exists : introduction via : exists ... elimination via : destruct ...

Lemma example_exists : x : nat, x = 0.
Proof.
  0. reflexivity.
Qed.

  • assumption if the current goal is exactly one of the hypothesis (cf. the axiom rule in logic).
  • For all primitive connectors (quantification , implication ):
    • introduction via intro (or one of its variants intros,
    intro x, intros x y ...)
    • elimination via apply H (where H is the name of
    the hypothesis to eliminate).
  • The other connectors (which are actually inductive definitions) may ideally be introduced by constructor and eliminated by destruct H. But the introduction frequently requires more ad-hoc tactics:
    • split for a conjunction (written /\)
    • left and right for a disjunction (written \/)
    • ... (written exists) for a quantification
    (where ... is the place where the existential witness is given)
    • No introduction tactic for False !
    • For True (seldom used in Coq), the introduction can
    be done via constructor, but nothing to eliminate.
  • Some abbreviations:
    • A negation ¬A (written ~A) is just a shortcut for AFalse.
    Hence introduction via intro (or intros a, giving a name forces the introduction) and elimination via apply H or destruct H (whether one wants to focus on the underlying or False).
    • An equivalence AB (written A<->B) is just a shortcut for (AB)/\(BA),
    hence is manipulated as a conjunction of implications.
  • Some automatic tactics : trivial, easy, auto, eauto, intuition, firstorder. See the Coq documentation for more details.
In today practical session, you will first try to prove the statements with no or little automatisation, and then experiment with these tactics in a second time. Indeed, it can be helpful later to know how to proceed steps-by-steps, since real-life proofs are seldom doable by just automatic tactics.
In Coq, a definition d could be replaced by its body thanks to the tactic unfold d.
The equality is handled by tactics:
  • reflexivity
  • symmetry
  • transitivity ... (where you need to give the intermediate
term) and
  • rewrite ... (give the equality name or lemma to use as a
rewrite rule) or
  • rewrite <- ... (right-to-left rewriting)
Inductive reasoning will be studied at length in coming lectures. Just to state it immediately:
  • induction ... (to do inductive reasoning, to be investigated in the next lectures).

This page has been generated by coqdoc