Library DProp.Prop
Require Import String.
Require Import DProp.Tactics.
Require Import List.
Import List.ListNotations.
Require Import DProp.Tactics.
Require Import List.
Import List.ListNotations.
Inductive sentence : Set :=
| p_top : sentence
| p_var : nat -> sentence
| p_conj : sentence -> sentence -> sentence
| p_neg : sentence -> sentence.
| p_top : sentence
| p_var : nat -> sentence
| p_conj : sentence -> sentence -> sentence
| p_neg : sentence -> sentence.
We ask Coq to infer that equality of sentences is decidable
Scheme Equality for sentence.
Check sentence_eq_dec. (* : : forall x y : sentence, {x = y} + {x <> y} *)
Check sentence_eq_dec. (* : : forall x y : sentence, {x = y} + {x <> y} *)
We'll give names to a few of the variables for convenience...
(* Notation " n " := (p_var n). *)
Notation "'P'" := (p_var 0).
Notation "'Q'" := (p_var 1).
Notation "'R'" := (p_var 2).
Notation "'P'" := (p_var 0).
Notation "'Q'" := (p_var 1).
Notation "'R'" := (p_var 2).
... and set up some pleasant notations. Notice some of these
notations define new connectives, like disjunction, in terms of
simpler ones.
These notations use Unicode. Proof General and CoqIDE both support
quick Unicode entry by typing TeX-like strings, see
- Coqide information (See "Using Unicode symbols")
Notation "⊤" := (p_top). (* \top *)
Notation "x ∧ y" := (p_conj x y) (at level 50). (* \land or \wedge *)
Notation "¬ x" := (p_neg x) (at level 10). (* \neg *)
Notation "⊥" := (¬ ⊤). (* \bot *)
Notation "x ∨ y" := (¬ (¬ x ∧ ¬ y)) (at level 50). (* \lor or \vee *)
Notation "x ⇒ y" := (¬ x ∨ y) (at level 70). (* \Rightarrow *)
Notation "x ∧ y" := (p_conj x y) (at level 50). (* \land or \wedge *)
Notation "¬ x" := (p_neg x) (at level 10). (* \neg *)
Notation "⊥" := (¬ ⊤). (* \bot *)
Notation "x ∨ y" := (¬ (¬ x ∧ ¬ y)) (at level 50). (* \lor or \vee *)
Notation "x ⇒ y" := (¬ x ∨ y) (at level 70). (* \Rightarrow *)
Next we define the semantics of our language. We want to translate
sentences into Coq-level propositions, i.e. elements of type
Prop. You might expect to use bool here instead. Both choices
"work," but using Prop is more general and more natural for the
metatheory we build here today, and ultimately it better captures
the idea that a semantics maps object-level propositions into
host-level propositions. Further information about this topic is
provided below.
A valuation might also be called an interpretation or a model,
especially in the context of higher-order logics.
Notice that we require a valuation to come with a proof that the
propositional symbols of our language are mapped to propositions
for which the excluded middle is provable (or taken as an
axiom). This is necessary for our logic to be sound, which is
discussed in the section on proof theory.
This coercion tells Coq that we can use a valuation as a
function nat -> Prop by using its val field, for convenience.
Model theory
Fixpoint denotation (v : valuation) (s : sentence) : Prop :=
match s with
| p_top => True
| p_var x => val x
| p_conj s1 s2 => denotation v s1 /\ denotation v s2
| p_neg s1 => not (denotation v s1)
end.
match s with
| p_top => True
| p_var x => val x
| p_conj s1 s2 => denotation v s1 /\ denotation v s2
| p_neg s1 => not (denotation v s1)
end.
Hereafter we let denotation take the valuation as an implicit
argument. This follows common mathematical practice in which one
generally assumes some valuation is clear from the context.
The law of excluded middle extends to the entire language straightforwardly.
Lemma full_lem : forall {v : valuation}, forall ϕ, ⟦ ϕ ⟧ \/ ⟦ ¬ ϕ ⟧.
Proof.
intros v ϕ. induction ϕ.
- simpl. intuition.
- exact (excluded_middle n).
- simpl. intuition.
- simpl. intuition.
Qed.
Proof.
intros v ϕ. induction ϕ.
- simpl. intuition.
- exact (excluded_middle n).
- simpl. intuition.
- simpl. intuition.
Qed.
This tactic splits a (Coq-level) proof into two cases: one in
which ϕ is true and one in which it is false.
Sample usage of the lem tactic.
Goal forall ϕ, forall v : valuation, ⟦ ϕ ⟧ \/ ⟦ ¬ ϕ ⟧.
Proof.
intros.
lem ϕ.
- left. auto.
- right. auto.
Qed.
Proof.
intros.
lem ϕ.
- left. auto.
- right. auto.
Qed.
In this file, we're not very interested in individual valuations. Instead,
we would like to define a notion of "truth" which abstracts over the valuation,
capturing the idea that a sentence is true under every possible interpretation.
The universal quantifier---the dependent product---can be thought
of as a kind of infinite conjunction asserting the truth of ϕ
under every interpretation.
But we aren't just interested in tautologies. Often we want to
restrict our attention to only those valuations which make certain
pre-chosen sentences true. We develop this idea now.
A set of sentences is satisfied or modeled by a valuation if
all of its elements are true under that valuation. Elements of Γ
may also be called "axioms."
A valuation that models Γ is called a model, unsurprisingly.
A set of axioms entails a sentence when that sentence is true
under all valuations that also satisfy the axioms, i.e. all
models. This relation also known as semantic consequence or
logical consequence.
We might also read this definition as a restricted variant of
tautology that intersects only over that subset of valuations
which satisfy Γ. This property is sometimes known as being /valid/
for Γ (but this word is sometimes used differently). When we
formalize our proof system, the most ideal outcome is that the
provable formulas are precisely the valid ones.
Definition entails (Γ : list sentence) (ϕ : sentence) :=
forall v, models v Γ -> ⟦ ϕ ⟧.
Notation "Γ ⊧ ϕ" := (entails Γ ϕ) (at level 70).
Hint Unfold tautology : dp.
Hint Unfold models : dp.
Hint Unfold entails : dp.
forall v, models v Γ -> ⟦ ϕ ⟧.
Notation "Γ ⊧ ϕ" := (entails Γ ϕ) (at level 70).
Hint Unfold tautology : dp.
Hint Unfold models : dp.
Hint Unfold entails : dp.
Use the following tactic to quickly get rid unfold many definitions.
Tactic Notation "unf" := (repeat autounfold with dp in *).
Exercise 1 (easy): Prove that a sentence is a tautology if and
only if it is entailed by the empty set of axioms.
Note: This is almost true by definition, but formalizing logics in
Coq often means proving such "obvious" facts.
Exercise 2 (easy): Prove the following tautology.
Exercise 3 (medium): Have we formalized implication correctly?
Check our work by proving if x ⇒ y is true and x is true, then
y is true.
Hint: Use the lem tactic to do case analysis on Q. You may
want to Unset Printing Notations.
Logically equivalent formulas are those which have the same
denotation under every interpretation. Since we are interpreting
into Prop, rather than bool, we define this notion using
biconditionality (at the Coq level) rather than equality of
boolean values, but the idea is the same.
Exercise 4 (easy): Prove that P and P ∧ ⊤ are semantically
equivalent.
Exercise 5 (medium): Prove that P and P ∨ ¬ ⊤ are
semantically equivalent. Use lem P.
Exercise 6 (hard): Why did the last exercise require law of
excluded middle, when Coq trivially proves the following
intuitionistic tautology? Can you prove (with pen and paper, not
in Coq) that the last exercise cannot be completed without LEM?
Proof theory
Reserved Notation "Γ ⊢ ϕ" (at level 90).
Inductive provable (Γ : list sentence) : sentence -> Prop :=
| j_var : forall ϕ, List.In ϕ Γ -> Γ ⊢ ϕ
| j_conj_intro : forall ϕ ψ, Γ ⊢ ϕ -> Γ ⊢ ψ -> Γ ⊢ ϕ ∧ ψ
| j_conj_elim1 : forall ϕ ψ, Γ ⊢ ϕ ∧ ψ -> Γ ⊢ ϕ
| j_conj_elim2 : forall ϕ ψ, Γ ⊢ ϕ ∧ ψ -> Γ ⊢ ψ
| j_neg_intro : forall ϕ, ϕ :: Γ ⊢ ⊥ -> Γ ⊢ ¬ ϕ
| j_neg_elim : forall ϕ ψ, Γ ⊢ ϕ -> Γ ⊢ ¬ ϕ -> Γ ⊢ ψ
| j_raa: forall ϕ, (¬ ϕ :: Γ ⊢ ¬ ⊤) -> Γ ⊢ ϕ
| j_top_intro : Γ ⊢ ⊤
where "Γ ⊢ ϕ" := (provable Γ ϕ).
The next two lemmas are fundamental properties of most logics.
Substitution shows that provable hypotheses are
redundant. Weakening says unused hypotheses are okay and plays an
important role in proving substitution.
Lemma weakening : forall Γ1 Γ2 Γ3 ϕ,
Γ1 ++ Γ2 ⊢ ϕ ->
Γ1 ++ Γ3 ++ Γ2 ⊢ ϕ.
Proof.
intros ? ? ? ? Hϕ.
remember (Γ1 ++ Γ2) as Γ.
generalize dependent Γ1.
induction Hϕ; intros; subst.
- apply j_var.
rewrite ?List.in_app_iff in *.
intuition.
- apply j_conj_intro; auto.
- eapply j_conj_elim1. eauto.
- eapply j_conj_elim2. eauto.
- eapply j_neg_intro.
specialize (IHHϕ (ϕ :: Γ1)). simpl in *. auto.
- eapply j_neg_elim. eapply IHHϕ1; reflexivity. auto.
- eapply j_raa. specialize (IHHϕ (¬ ϕ :: Γ1)).
simpl in *. eauto.
- apply j_top_intro.
Qed.
Γ1 ++ Γ2 ⊢ ϕ ->
Γ1 ++ Γ3 ++ Γ2 ⊢ ϕ.
Proof.
intros ? ? ? ? Hϕ.
remember (Γ1 ++ Γ2) as Γ.
generalize dependent Γ1.
induction Hϕ; intros; subst.
- apply j_var.
rewrite ?List.in_app_iff in *.
intuition.
- apply j_conj_intro; auto.
- eapply j_conj_elim1. eauto.
- eapply j_conj_elim2. eauto.
- eapply j_neg_intro.
specialize (IHHϕ (ϕ :: Γ1)). simpl in *. auto.
- eapply j_neg_elim. eapply IHHϕ1; reflexivity. auto.
- eapply j_raa. specialize (IHHϕ (¬ ϕ :: Γ1)).
simpl in *. eauto.
- apply j_top_intro.
Qed.
To picture substitution, imagine a natural deduction derivation D
that uses of ψ using hypotheses from Γ1, Γ2, and ϕ. Given a
natural deduction derivation E of ϕ, picture "moving up" the tree
of D and finding all leaves that introduce the axiom ϕ, and
replace them with E. Through Curry-Howard, this is essentially
function evaluation.
Lemma substitution : forall Γ1 ϕ Γ2 ψ,
Γ1 ++ (ϕ :: Γ2) ⊢ ψ ->
Γ2 ⊢ ϕ ->
Γ1 ++ Γ2 ⊢ ψ.
Proof.
intros ? ? ? ? J1 J2.
remember (List.app Γ1 (ϕ :: Γ2)).
generalize dependent Γ1.
induction J1; intros Γ1 Eq; subst.
- destruct (sentence_eq_dec ϕ ϕ0).
{ (* Equal *)
subst.
replace (Γ1 ++ Γ2) with (nil ++ Γ1 ++ Γ2) by (rewrite app_nil_l; auto).
apply weakening. rewrite app_nil_l. auto. }
{ (* Unequal *)
apply j_var.
rewrite ?List.in_app_iff in *.
destruct H; try tauto.
inversion H; subst. contradiction. tauto. }
- apply j_conj_intro; auto.
- eapply j_conj_elim1; auto.
- eapply j_conj_elim2; auto.
- apply j_neg_intro. specialize (IHJ1 (ϕ0 :: Γ1)). simpl in *. eauto.
- apply j_neg_elim with (ϕ := ϕ0); auto.
- apply j_raa.
specialize (IHJ1 (¬ ϕ0 :: Γ1)). simpl in *. eauto.
- apply j_top_intro.
Qed.
Γ1 ++ (ϕ :: Γ2) ⊢ ψ ->
Γ2 ⊢ ϕ ->
Γ1 ++ Γ2 ⊢ ψ.
Proof.
intros ? ? ? ? J1 J2.
remember (List.app Γ1 (ϕ :: Γ2)).
generalize dependent Γ1.
induction J1; intros Γ1 Eq; subst.
- destruct (sentence_eq_dec ϕ ϕ0).
{ (* Equal *)
subst.
replace (Γ1 ++ Γ2) with (nil ++ Γ1 ++ Γ2) by (rewrite app_nil_l; auto).
apply weakening. rewrite app_nil_l. auto. }
{ (* Unequal *)
apply j_var.
rewrite ?List.in_app_iff in *.
destruct H; try tauto.
inversion H; subst. contradiction. tauto. }
- apply j_conj_intro; auto.
- eapply j_conj_elim1; auto.
- eapply j_conj_elim2; auto.
- apply j_neg_intro. specialize (IHJ1 (ϕ0 :: Γ1)). simpl in *. eauto.
- apply j_neg_elim with (ϕ := ϕ0); auto.
- apply j_raa.
specialize (IHJ1 (¬ ϕ0 :: Γ1)). simpl in *. eauto.
- apply j_top_intro.
Qed.
Relating proofs to truths
- a "proofs" of a "proposition"
- a model of the axioms Γ
Our next theorem is the converse of soundness, known as semantic
completeness. It states that logically valid sentences have
proofs. There is another notion of completeness ("syntactical
(in)completeness") discussed later. It is sometimes clear from
context which notion of completeness one has in mind, since
well-known logics have equally well-known completeness properties,
but this can also be a source of confusion to newcomers. The
reader is advised to be very careful when reading discussions of
this topic on the internet, which are frequently confusing and
often simply incorrect.
Many logics do not have semantic completeness. Fortunately,
propositional logic is one of the logics with this
property. Unfortunately, proving this is difficult, even for such
a simple system. In fact, that we won't even try to prove it in
Coq. You are encouraged to attempt it to find where the difficulty
lies.
Why should this be hard to prove? Consider the premise: In all
valuations satisfying every sentence of Γ, ϕ is true. This quite
abstract condition is a statement about truth tables. If we read
the statement of completeness as the type of a constructive proof
in Coq, we find a challenging problem: "Given the fact that all
rows of a truth table that happen to satisfy each formula in Γ
also satisfy ϕ, find a natural deduction derivation of ϕ."
Notice the premise gives us almost no information about ϕ, not
even its basic structure. The sentences of Γ may also be very
complex, perhaps more complex than ϕ, so they don't necessary tell
us anything about ϕ's subformulas. In fact, while ϕ may be true
under every Γ-satisfying valuation, each Γ-satisfying row of the
truth table might make ϕ true "for different reasons," so to speak
(try some examples). Where do we get started?
Nonetheless, this theorem is true, and can even be proved
constructively. For now, we admit it without proof.
Theorem completeness : forall Γ ϕ, Γ ⊧ ϕ -> Γ ⊢ ϕ.
Proof.
intros Γ ϕ.
intros mod.
unf.
(* ????????????????????????????? *)
Admitted.
Proof.
intros Γ ϕ.
intros mod.
unf.
(* ????????????????????????????? *)
Admitted.
Exercise 8 (easy? hard?): Attempt to prove the following: "There
is a sentence of propositional logic that is neither provable nor
disprovable from the empty set of axioms."
Use the formula P (a simple propositional variable) as ϕ. Attempt
to prove this by induction on the derivation of P. A template has been provided.
What happens?
Theorem syntactical_incompleteness : exists ϕ, not (nil ⊢ ϕ) /\ not (nil ⊢ ¬ ϕ).
Proof.
exists P.
split; intro J.
- Case "P is not provable".
Proof.
exists P.
split; intro J.
- Case "P is not provable".
We want to scrutinize the supposed derivation of P.
If we just do induction J, Coq will automatically generalize over the context and conclusion P,
as if were trying to show there are no proofs of anything. To avoid this silly behavior,
we tell Coq to remember both the conclusion P and the emptiness of the context.
We also discharge trivial cases instantly.
Ltac cleanup := repeat match goal with | H : _ |- _ => specialize (H eq_refl) end.
remember nil as empty; remember P as conc;
induction J; subst; cleanup.
+ SCase "j_var".
(* for you to finish *)
admit.
+ SCase "j_conj_intro".
(* contradiction, the conclusions don't match *)
inversion Heqconc.
+ SCase "j_conj_elim1".
(* Why doesn't the induction hypothesis apply?
What happens if you try inversion on the J? Why?
*)
admit.
+ SCase "j_conj_elim2".
(* Same story as above. *)
admit.
+ SCase "j_neg_intro".
(* contradiction, the conclusions don't match *)
inversion Heqconc.
+ SCase "j_neg_elim".
(* Why doesn't the induction hypothesis apply?
What happens if you try inversion on the J? Why?
*)
admit.
+ SCase "j_raa".
(* Why doesn't the induction hypothesis apply?
What happens if you try inversion on the J? Why?
*)
admit.
+ SCase "j_top_intro".
(* contradiction, the conclusions don't match *)
inversion Heqconc.
- Case "~ P is not provable".
remember nil as empty; remember (¬ P) as conc;
induction J; subst; cleanup.
+ SCase "j_var".
(* for you to finish *)
admit.
+ SCase "j_conj_intro".
(* contradiction, the conclusions don't match *)
inversion Heqconc.
+ SCase "j_conj_elim1".
(* What happens if you try inversion on the J? Why?
*)
admit.
+ SCase "j_conj_elim2".
(* Same story as above. *)
admit.
+ SCase "j_neg_intro".
inversion Heqconc. subst; clear Heqconc.
(* What happens if you try inversion on the J? Why? *)
admit.
+ SCase "j_neg_elim".
(* What happens if you try inversion on the J? Why?
*)
admit.
+ SCase "j_raa".
(* What happens if you try inversion on the J? Why?
*)
admit.
+ SCase "j_top_intro".
(* contradiction, the conclusions don't match *)
inversion Heqconc.
Abort.
remember nil as empty; remember P as conc;
induction J; subst; cleanup.
+ SCase "j_var".
(* for you to finish *)
admit.
+ SCase "j_conj_intro".
(* contradiction, the conclusions don't match *)
inversion Heqconc.
+ SCase "j_conj_elim1".
(* Why doesn't the induction hypothesis apply?
What happens if you try inversion on the J? Why?
*)
admit.
+ SCase "j_conj_elim2".
(* Same story as above. *)
admit.
+ SCase "j_neg_intro".
(* contradiction, the conclusions don't match *)
inversion Heqconc.
+ SCase "j_neg_elim".
(* Why doesn't the induction hypothesis apply?
What happens if you try inversion on the J? Why?
*)
admit.
+ SCase "j_raa".
(* Why doesn't the induction hypothesis apply?
What happens if you try inversion on the J? Why?
*)
admit.
+ SCase "j_top_intro".
(* contradiction, the conclusions don't match *)
inversion Heqconc.
- Case "~ P is not provable".
remember nil as empty; remember (¬ P) as conc;
induction J; subst; cleanup.
+ SCase "j_var".
(* for you to finish *)
admit.
+ SCase "j_conj_intro".
(* contradiction, the conclusions don't match *)
inversion Heqconc.
+ SCase "j_conj_elim1".
(* What happens if you try inversion on the J? Why?
*)
admit.
+ SCase "j_conj_elim2".
(* Same story as above. *)
admit.
+ SCase "j_neg_intro".
inversion Heqconc. subst; clear Heqconc.
(* What happens if you try inversion on the J? Why? *)
admit.
+ SCase "j_neg_elim".
(* What happens if you try inversion on the J? Why?
*)
admit.
+ SCase "j_raa".
(* What happens if you try inversion on the J? Why?
*)
admit.
+ SCase "j_top_intro".
(* contradiction, the conclusions don't match *)
inversion Heqconc.
Abort.
Exercise 9 (easy):
What do the hard cases above have in common? (Hint: Look at the
names of the rules that caused us difficulty.).
The theorem above seems obvious, but it is hard to show by
examining proof trees. The trouble is that in each elimination
rule, we cannot be confident that the rule used to prove the
hypothesis is a corresponding introduction rule. Indeed, it might not
be: Perhaps we prove P by proving P /\ Q, and perhaps to prove
/that/ we use reductio ad absurdum (for example). Can you
rule this out?
Although implication is not defined primitively, it is an
illustrative example: Perhaps we prove P by showing Q ⇒ P and
Q. How can we rule this out? Fundamentally we need a
cut-elimination theorem, specifically the corollary of a
canonical form lemma (or inversion lemma). This would justify a
more controlled case analysis by showing, without loss of
generality, that we may assume proofs are of a particularly simple
form.
We won't prove this hard theorem here. Instead, we can show the
above theorem by using soundness: We cannot prove P because it
is not true in all models!
Definition vtrue : valuation :=
{| val := fun n => True ; excluded_middle := ltac:(intuition) |}.
Definition vfalse : valuation :=
{| val := fun n => False ; excluded_middle := ltac:(intuition) |}.
Exercise 10 (medium): Prove the following statement using
the two constant valuations defined above.
Commentary
Understanding Incompleteness
- True, in the sense that it is true when interpreted into the model given by nat.
- Unprovable: There is no natural deduction derivation of this sentence.
- invalid: There are models in Coq of first-order Peano arithmetic in which the sentence is not true.
Why is incompleteness important? Or is it?
- Coq is incomplete because it neither proves nor disproves LEM
- Group theory is incomplete because it neither proves nor disproves the commutativity property
- Propositional logic is incomplete because it neither proves nor disproves any atomic propositional symbol P
- we can recognize valid proofs when we see them
- the theory can prove basic facts (things which can be proved without even using induction)