(** * Logic: Logic in Coq *) (* Version of 9/13/2009 *) Require Export Props. (** Like its built-in programming language, Coq's built-in logic is extremely small: universal quantification ([forall]) and implication ([->]) are primitive, but all the other familiar logical connectives -- conjunction, disjunction, negation, existential quantification, even equality -- can be defined using just these and [Inductive]. *) (* ########################################################### *) (** * Conjunction *) (** The logical conjunction of propositions [A] and [B] is represented by the following inductively defined proposition. *) Inductive and (A B : Prop) : Prop := conj : A -> B -> (and A B). (** Note that, like the definition of [ev], this definition is parameterized; however, in this case, the parameters are themselves propositions. *) (** The intuition behind this definition is simple: to construct evidence for [and A B], we must provide evidence for [A] and evidence for [B]. More precisely: - [conj e1 e2] can be taken as evidence for [and A B] if [e1] is evidence for [A] and [e2] is evidence for [B]; and - this is the _only_ way to give evidence for [and A B] -- that is, if someone gives us evidence for [and A B], we know it must have the form [conj e1 e2], where [e1] is evidence for [A] and [e2] is evidence for [B]. *) (** Since we'll be using conjunction a lot, let's introduce a more familiar-looking infix notation for it. *) Notation "A /\ B" := (and A B) : type_scope. (** (The [type_scope] annotation tells Coq that this notation will be appearing in propositions, not values.) *) (** Consider the "type" of the constructor [conj] *) (* Check conj. *) (** Notice that it "takes 4 inputs", namely the propositions [A],[B] and evidence of [A], [B], and "returns as output", the evidence of [A /\ B]. *) (** **** Exercise: 1 star (and_ind_principle) *) (** See if you can predict the induction principle for conjunction. *) (* Check and_ind. *) (** [] *) (** Besides the elegance of building everything up from a tiny foundation, what's nice about defining conjunction this way is that we can prove statements involving conjunction using the tactics that we already know. For example, if the goal statement is a conjuction, we can prove it by applying the single constructor [conj], which (as can be seen from the type of [conj]) solves the current goal and leaves the two parts of the conjunction as subgoals to be proved separately. *) Theorem and_example : (ev 0) /\ (ev 4). Proof. apply conj. Case "left". apply ev_0. Case "right". apply ev_SS. apply ev_SS. apply ev_0. Qed. (** Let's take a look at the proof object for the above theorem. *) (* Print and_example. *) (** Note that the proof is of the form [[ conj (ev 0) (ev 4) (Pf of ev 0) (Pf of ev 4) ]] which is what you'd expect, given the type of [conj] *) (** The [split] tactic is a convenient shorthand for [apply conj]. *) Theorem and_example' : (ev 0) /\ (ev 4). Proof. split. Case "left". apply ev_0. Case "right". apply ev_SS. apply ev_SS. apply ev_0. Qed. (** Conversely, the [inversion] tactic can be used to investigate a conjunction hypothesis in the context and calculate what evidence must have been used to build it. *) Theorem proj1 : forall A B : Prop, A /\ B -> A. Proof. intros A B H. inversion H as [HA HB]. apply HA. Qed. (** **** Exercise: 1 star *) Theorem proj2 : forall A B : Prop, A /\ B -> B. Proof. (* FILL IN HERE *) Admitted. (** [] *) Theorem and_commut : forall A B : Prop, A /\ B -> B /\ A. Proof. intros A B H. split. Case "left". apply proj2 with (A:=A). apply H. Case "right". apply proj1 with (B:=B). apply H. Qed. (** **** Exercise: 2 stars *) (** Notice the nested inversion breaks [H : A /\ (B /\ C)] down into [HA: A], [HB : B] and [HC : C]. Finish the proof from there: *) Theorem and_assoc : forall A B C : Prop, A /\ (B /\ C) -> (A /\ B) /\ C. Proof. intros A B C H. inversion H as [HA [HB HC]]. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 2 stars *) (** Now we can prove the other direction of the equivalence of [even] and [ev]. Notice that the left-hand conjunct here is the statement we are actually interested in; the right-hand conjunct is needed in order to make the induction hypothesis strong enough that we can carry out the reasoning in the inductive step. (To see why this is needed, try proving the left conjunct by itself and observe where things get stuck.) *) Theorem even_ev : forall n : nat, (even n -> ev n) /\ (even (S n) -> ev (S n)). Proof. (* Hint: Use induction on [n]. *) (* FILL IN HERE *) Admitted. (** [] *) (* ###################################################### *) (** * Iff *) (** The familiar logical "if and only if" is just the conjunction of two implications. *) Definition iff (A B : Prop) := (A -> B) /\ (B -> A). Notation "A <-> B" := (iff A B) (at level 95, no associativity) : type_scope. Theorem iff_implies : forall A B : Prop, (A <-> B) -> A -> B. Proof. intros A B H. inversion H as [HAB HBA]. apply HAB. Qed. Theorem iff_sym : forall A B : Prop, (A <-> B) -> (B <-> A). Proof. intros A B H. inversion H as [HAB HBA]. split. (* Note that [split] is just a bit smarter than [apply conj], which would have required an [unfold iff] first. *) Case "->". apply HBA. Case "<-". apply HAB. Qed. (** **** Exercise: 1 star (iff_properties) *) (** Using the above proof that [<->] is symmetric ([iff_sym]) as a guide, prove that it is also reflexive and transitive. *) Theorem iff_refl : forall A : Prop, A <-> A. Proof. (* FILL IN HERE *) Admitted. (** Hint: If you have an iff hypothesis in the context, you can use [inversion] to break it into two separate implications. (Think about why this works.) *) Theorem iff_trans : forall A B C : Prop, (A <-> B) -> (B <-> C) -> (A <-> C). Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 2 stars *) (** We have seen that the families of propositions [MyProp] and [ev] actually characterize the same set of numbers (the even ones). Prove that [MyProp n <-> ev n] for all [n] ([MyProp_iff_ev] in [Logic.v]). Just for fun, write your proof as an explicit proof object, rather than using tactics. *) Definition MyProp_iff_ev : forall n, MyProp n <-> ev n := (* FILL IN HERE *) admit. (** [] *) (** Unfortunately, propositions phrased with [<->] are a bit inconvenient to use as hypotheses or lemmas, because they have to be deconstructed into their two directed components in order to be applied. (The basic problem is that there's no way to apply an iff proposition directly. If it's a hypothesis, you can invert it, which is tedious; if it's a lemma, you have to destruct it into hypotheses, which is worse.) Consequently, many Coq developments avoid [<->], despite its appealing compactness. It can actually be made much more convenient using a Coq feature called "setoid rewriting," but that is a bit beyond the scope of this course. *) (* ############################################################ *) (** * Disjunction *) (** Disjunction ("logical or") can also be defined as an inductive proposition. *) Inductive or (A B : Prop) : Prop := | or_introl : A -> or A B | or_intror : B -> or A B. Notation "A \/ B" := (or A B) : type_scope. (** **** Exercise: 1 star (or_ind_principle) *) (** See if you can predict the induction principle for disjunction. *) (* Check or_ind. *) (** [] *) (** Consider the "type" of the constructor [or_introl] *) (* Check or_introl. *) (** Notice that it "takes 3 inputs", namely the propositions [A],[B] and evidence of [A], and "returns as output", the evidence of [A /\ B]. Next, look at the "type" of [or_intror] *) (* Check or_intror. *) (** It is like [or_introl] but instead of evidence of [A], requires evidence of [B] *) (** Intuition: There are two ways of giving evidence for [A \/ B]: - give evidence for [A] (and say that it is [A] you are giving evidence for! -- this is the function of the [or_introl] constructor) - give evidence for [B], tagged with the [or_intror] constructor. *) (** Since [A \/ B] has two constructors, doing [inversion] on a hypothesis of type [A \/ B] yields two subgoals. *) Theorem or_commut : forall A B : Prop, A \/ B -> B \/ A. Proof. intros A B H. inversion H as [HA | HB]. Case "left". apply or_intror. apply HA. Case "right". apply or_introl. apply HB. Qed. (** From here on, we'll use the handy tactics [left] and [right] in place of [apply or_introl] and [apply or_intror]. *) Theorem or_commut' : forall A B : Prop, A \/ B -> B \/ A. Proof. intros A B H. inversion H as [HA | HB]. Case "left". right. apply HA. Case "right". left. apply HB. Qed. Theorem or_distributes_over_and_1 : forall A B C : Prop, A \/ (B /\ C) -> (A \/ B) /\ (A \/ C). Proof. intros A B C. intros H. inversion H as [HA | [HB HC]]. Case "left". split. SCase "left". left. apply HA. SCase "right". left. apply HA. Case "right". split. SCase "left". right. apply HB. SCase "right". right. apply HC. Qed. (** **** Exercise: 2 stars *) Theorem or_distributes_over_and_2 : forall A B C : Prop, (A \/ B) /\ (A \/ C) -> A \/ (B /\ C). Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 1 star *) Theorem or_distributes_over_and : forall A B C : Prop, A \/ (B /\ C) <-> (A \/ B) /\ (A \/ C). Proof. (* FILL IN HERE *) Admitted. (** [] *) (* ################################################### *) (** * Relating /\ and \/ with andb and orb *) (** We've already seen several places where analogous structures can be found in Coq's computational ([Type]) and logical ([Prop]) worlds. Here is one more: the boolean operators [andb] and [orb] are obviously analogs, in some sense, of the logical connectives [/\] and [\/]. This analogy can be made more precise by the following theorems, which show how to "translate" knowledge about [andb] and [orb]'s behaviors on certain inputs into propositional facts about those inputs. *) Theorem andb_true : forall b c, andb b c = true -> b = true /\ c = true. Proof. intros b c H. destruct b. destruct c. apply conj. reflexivity. reflexivity. inversion H. inversion H. Qed. (** **** Exercise: 1 star (bool_prop) *) (** Prove these properties relating operations on booleans and logical connectives in [Prop]. *) Theorem andb_false : forall b c, andb b c = false -> b = false \/ c = false. Proof. (* FILL IN HERE *) Admitted. Theorem orb_true : forall b c, orb b c = true -> b = true \/ c = true. Proof. (* FILL IN HERE *) Admitted. Theorem orb_false : forall b c, orb b c = false -> b = false /\ c = false. Proof. (* FILL IN HERE *) Admitted. (** [] *) (* ################################################### *) (** * Falsehood *) (** Logical falsehood can be represented in Coq as an inductively defined proposition with no constructors. *) Inductive False : Prop := . (** Intuition: [False] is a proposition for which there is no way to give evidence. *) (** **** Exercise: 1 star (False_ind_principle) *) (** Can you predict the induction principle for falsehood? *) (* Check False_ind. *) (** [] *) (** Since [False] has no constructors, inverting it always yields zero subgoals, allowing us to immediately prove any goal. *) Theorem False_implies_nonsense : False -> plus 2 2 = 5. Proof. intros contra. inversion contra. Qed. (** What happened? [inversion] breaks [contra] into each of its possible cases, and yields a subgoal for each case. As [contra] is evidence for [False], it has _no_ possible cases, hence, there are no possible subgoals and the proof is done. *) (** Conversely, the only way to prove [False] is if there is already something nonsensical or contradictory in the context: *) Theorem nonsense_implies_False : plus 2 2 = 5 -> False. Proof. intros contra. inversion contra. Qed. (** Actually, since the proof of [False_implies_nonsense] doesn't actually have anything to do with the specific nonsensical thing being proved; it can easily be generalized to work for an arbitrary [P]: *) Theorem ex_falso_quodlibet : forall (P:Prop), False -> P. Proof. intros P contra. inversion contra. Qed. (** The Latin _ex falso quodlibet_ means, literally, "from falsehood follows whatever you please." This theorem is also known as the _principle of explosion_. *) (* #################################################### *) (** * Truth *) (** Since we have defined falsehood in Coq, we should also mention that it is, of course, possible to define truth in the same way. *) (** **** Exercise: 2 stars *) (** Define [True] as another inductively defined proposition. What induction principle will Coq generate for your definition? (The intution is that [True] should be a proposition for which it is trivial to give evidence. Alternatively, you may find it easiest to start with the induction principle and work backwards to the inductive definition.) *) (* FILL IN HERE *) (** [] *) (** However, unlike [False], which we'll use extensively, [True] is basically a theoretical curiosity: since it is trivial to prove as a goal, it carries no useful information as a hypothesis. *) (* #################################################### *) (** * Negation *) (** The logical complement of a proposition [A] is written [not A] or, for shorthand, [~A]: *) Definition not (A:Prop) := A -> False. (** The intuition is that, if [A] is not true, then anything at all (even [False]) should follow from assuming [A]. *) Notation "~ x" := (not x) : type_scope. (* Check not. *) (** It takes a little practice to get used to working with negation in Coq. Even though you can see perfectly well why something is true, it can be a little hard at first to figure out how to get things into the right configuration so that Coq can see it! Here are proofs of a view familiar facts about negation to get you warmed up. *) Theorem not_False : ~ False. Proof. unfold not. intros H. inversion H. Qed. Theorem contradiction_implies_anything : forall A B : Prop, (A /\ ~A) -> B. Proof. intros A B H. inversion H as [HA HNA]. unfold not in HNA. apply HNA in HA. inversion HA. Qed. Theorem double_neg : forall A : Prop, A -> ~~A. Proof. intros A H. unfold not. intros G. apply G. apply H. Qed. (** **** Exercise: 2 stars (double_neg_inf) *) (** As an exercise, write an informal proof of [double_neg]: Theorem: For any prop [A], [A -> ~~A] Proof: (* FILL IN HERE *) [] *) (** **** Exercise: 2 stars *) Theorem contrapositive : forall A B : Prop, (A -> B) -> (~B -> ~A). Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 1 star *) Theorem not_both_true_and_false : forall A : Prop, ~ (A /\ ~A). Proof. (* FILL IN HERE *) Admitted. (** [] *) Theorem five_not_even : ~ ev 5. Proof. unfold not. intros Hev5. inversion Hev5 as [|n Hev3 Heqn]. inversion Hev3 as [|n' Hev1 Heqn']. inversion Hev1. Qed. (** **** Exercise: 1 star *) (** Theorem [five_not_even] in [Logic.v] confirms the unsurprising fact that that five is not an even number. Prove this more interesting fact: *) Theorem ev_not_ev_S : forall n, ev n -> ~ ev (S n). Proof. unfold not. intros n H. induction H. (* not n! *) (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 4 stars, optional (classical_axioms) *) (** For those who like a challenge, here is an exercise taken from the Coq'Art book (p. 123). The following five statements are often considered as characterizations of classical logic (as opposed to constructive logic, which is what is "built in" to Coq). We can't prove them in Coq, but we can consistently add any one of them as an unproven axiom if we wish to work in classical logic. Prove that these five propositions are equivalent. *) Definition peirce := forall P Q: Prop, ((P->Q)->P)->P. Definition classic := forall P:Prop, ~~P -> P. Definition excluded_middle := forall P:Prop, P \/~P. Definition de_morgan_not_and_not := forall P Q:Prop, ~(~P/\~Q) -> P\/Q. Definition implies_to_or := forall P Q:Prop, (P->Q) -> (~P\/Q). (* FILL IN HERE *) (** [] *) (* ########################################################## *) (** * Inequality *) (** Saying [x <> y] is just the same as saying [~(x = y)]. *) Notation "x <> y" := (~ (x = y)) : type_scope. (** Since inequality involves a negation, it again requires a little practice to be able to work with it fluently. Here is one very useful trick. If you are trying to prove a goal that is nonsensical (e.g., the goal state is [false = true]), apply the lemma [ex_falso_quodlibet] to change the goal to [False]. This makes it easier to use assumptions of the form [~P] that are available in the context -- in particular, assumptions of the form [x<>y]. *) Theorem not_false_then_true : forall b : bool, b <> false -> b = true. Proof. intros b H. destruct b. Case "b = true". reflexivity. Case "b = false". unfold not in H. (* optional step *) apply ex_falso_quodlibet. apply H. reflexivity. Qed. (** **** Exercise: 2 stars *) Theorem not_eq_beq_false : forall n n' : nat, n <> n' -> beq_nat n n' = false. Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 4 stars, optional *) Theorem beq_false_not_eq : forall n m, false = beq_nat n m -> n <> m. Proof. (* FILL IN HERE *) Admitted. (** [] *) (* ############################################################ *) (** * Existential quantification *) (** Another extremely important logical connective is _existential quantification_. *) Inductive ex (X : Type) (P : X -> Prop) : Prop := ex_intro : forall witness:X, P witness -> ex X P. (* Check ex_intro. *) (** The intuition behind this definition is that, in order to give evidence for the assertion "there exists an [x] for which the proposition [P] holds" we must actually name a _witness_ -- a specific value [x] -- and then give evidence for [P x]. *) (** We can use Coq's notation definition facility to introduce more standard notation for writing existentially quantified propositions, exactly parallel to the built-in syntax for universally quantified propositions. Instead of writing [(ex nat (fun x => ev x))] to express the proposition that there exists some number that is even, for example, we can write [exists x:nat, ev x]. (It is not necessary to understand the details of how the definition works.) *) Notation "'exists' x , p" := (ex _ (fun x => p)) (at level 200, x ident, right associativity) : type_scope. Notation "'exists' x : X , p" := (ex _ (fun x:X => p)) (at level 200, x ident, right associativity) : type_scope. (** We can use the same tactics as always for manipulate existentials. For example, if to prove an existential, we [apply] the constructor [ex_intro]. Since the premise of [ex_intro] involves a variable ([witness]) that does not appear in its conclusion, we need to explicitly give its value when we use [apply]. *) Example exists_example_1 : exists n, plus n (mult n n) = 6. Proof. apply ex_intro with (witness:=2). reflexivity. Qed. (** Note that we have to explicitly give the witness. *) (** Or, instead of writing [apply ex_intro with e], we can write the convenient shorthand [exists e]. *) Example exists_example_1' : exists n, plus n (mult n n) = 6. Proof. exists 2. reflexivity. Qed. (** Conversely, if we have an existential hypothesis in the context, we can eliminate it with [inversion]. Note the use of the [as...] pattern to name the variable that Coq introduces to name the witness value and get evidence that the hypothesis holds for the witness. (If we don't explicitly choose one, Coq will just call it [witness], which makes proofs confusing.) *) Theorem exists_example_2 : forall n, (exists m, n = plus 4 m) -> (exists o, n = plus 2 o). Proof. intros n H. inversion H as [m Hm]. exists (plus 2 m). apply Hm. Qed. (** **** Exercise: 1 star *) (** Prove that "[P] holds for all [x]" and "there is no [x] for which [P] does not hold" are equivalent assertions. *) Theorem dist_not_exists : forall (X:Type) (P : X -> Prop), (forall x, P x) -> ~ (exists x, ~ P x). Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 2 stars *) (** Prove that existential quantification distributes over disjunction. *) Theorem dist_exists_or : forall (X:Type) (P Q : X -> Prop), (exists x, P x \/ Q x) <-> (exists x, P x) \/ (exists x, Q x). Proof. (* FILL IN HERE *) Admitted. (** [] *) (* ###################################################### *) (** * Equality. *) (** Even Coq's equality relation is not actually built in. It has the following inductive definition. (We enclose this definition in a module to avoid confusion with the standard library equality (which we have used extensively already). *) Module MyEquality. Inductive eq (A:Type) : A -> A -> Prop := refl_equal : forall x, eq A x x. Notation "x = y" := (eq _ x y) (at level 70, no associativity) : type_scope. (** This definition is a bit subtle. The way to think about it is that, given a set [A], it defines a _family_ of propositions "[x] is equal to [y]," indexed by pairs of values ([x] and [y]) from [A]. There is just one way of constructing evidence for members of this family: by applying the constructor [refl_equal] to two identical arguments. *) (** Here is an alternate definition, from the Coq standard library. *) Inductive eq' (A:Type) (x:A) : A -> Prop := refl_equal' : eq' A x x. Notation "x =' y" := (eq' _ x y) (at level 70, no associativity) : type_scope. (** **** Exercise: 3 stars, optional *) (** Verify that the two definitions of equality are equivalent. *) Theorem two_defs_of_eq_coincide : forall (A:Type) x y, eq A x y <-> eq' A x y. Proof. (* FILL IN HERE *) Admitted. (** [] *) (** The advantage of the second definition is that the induction principle that Coq derives for it is precisely the familiar principle of _Leibniz equality_: what we mean when we say "[x] and [y] are equal" is that every property on [A] that is true of [x] is also true of [y]. *) (* Check eq'_ind. *) End MyEquality. (* ####################################################### *) (** * Inversion, again *) (** We've seen [inversion] used with both equality hypotheses and hypotheses asserting inductively defined propositions. Now that we've seen that these are actually the same thing, let's take a closer look at how [inversion] behaves... In general, the [inversion] tactic - takes a hypothesis [H] whose type is some inductively defined proposition [P] - for each constructor [C] in [P]'s definition - generates a new subgoal in which we assume [H] was built with [C] - adds the arguments (premises) of [C] to the context of the subgoal as extra hypotheses - matches the conclusion (result type) of [C] against the current goal and calculates a set of equalities that must hold in order for [C] to be applicable - adds these equalities to the context of the subgoal - if the equalities are not satisfiable (e.g., they involve things like [S n = O], immediately solves the subgoal. EXAMPLE: If we invert a hypothesis built with [or], there are two constructors, so two subgoals get generated. The conclusion (result type) of the constructor -- [A \/ B] -- doesn't place any restrictions on the form of [A] or [B], so we don't get any extra equalities in the context of the subgoal. EXAMPLE: If we invert a hypothesis built with [and], there is only one constructor, so only one subgoal gets generated. Again, the conclusion (result type) of the constructor -- [A /\ B] -- doesn't place any restrictions on the form of [A] or [B], so we don't get any extra equalities in the context of the subgoal. The constructor does have two arguments, though, and these can be seen in the context in the subgoal. EXAMPLE: If we invert a hypothesis built with [eq], there is again only one constructor, so only one subgoal gets generated. Now, though, the form of the [refl_equal] constructor does give us some extra information: it tells us that the two arguments to [eq] must be equal! The [inversion] tactic adds this to the context. *) (* ####################################################### *) (** * Relations as propositions *) (** A proposition parameterized over a number (like [ev]) can be thought of as a _predicate_ -- i.e., it defines a subset of [nat], namely those numbers for which the proposition is provable. In the same way, a two-argument proposition can be thought of as a _relation_ -- i.e., it defines a set of pairs for which the proposition is provable. In the next few sections, we explore the consequences of this observation. *) Module FirstLe. (** We've already seen an inductive definition of one fundamental relation: equality. Another useful one is the "less than or equal to" relation on numbers: *) (** This definition should be fairly intuitive. It says that there are two ways to give evidence that one number is less than or equal to another: either observe that they are the same number, or give evidence that the first is less than or equal to the predecessor of the second. *) Inductive le : nat -> nat -> Prop := | le_n : forall n, le n n | le_S : forall n m, (le n m) -> (le n (S m)). (* Check le_ind. *) End FirstLe. (** This is a fine definition of the [<=] relation, but we can streamline it a little by observing that the left-hand argument [n] is the same everywhere in the definition, so we can actually make it a "general parameter" to the whole definition, rather than an argument to each constructor. *) Inductive le (n:nat) : nat -> Prop := | le_n : le n n | le_S : forall m, (le n m) -> (le n (S m)). Notation "m <= n" := (le m n). (** Why is the second one better (even though it looks less symmetric)? Because it gives us a simpler induction principle. *) (* Check le_ind. *) (** By contrast, the induction principle that Coq calculates for the first definition has a lot of extra quantifiers, which makes it messier to work with when proving things by induction. Here is the induction principle for the first [le]: *) (* Check le_ind. (the first one above) *) (* ===> le_ind : forall P : nat -> nat -> Prop, (forall n : nat, P n n) -> (forall n m : nat, FirstLe.le n m -> P n m -> P n (S m)) -> forall n n0 : nat, FirstLe.le n n0 -> P n n0 *) (** Proofs of facts about [<=] using the constructors [le_n] and [le_S] follow the same patterns as proofs about predicates, like [ev] in the previous chapter. We can [apply] the constructors to prove [<=] goals (e.g., to show that [3<=3] or [3<=6]), and we can use tactics like [inversion] to extract information from [<=] hypotheses in the context (e.g., to prove that [~(2 <= 1)].) *) (** Here are some sanity checks on the definition. (Notice that, although these are the same kind of simple "unit tests" as we gave for the testing functions we wrote in the first few lectures, we must construct their proofs explicitly -- [simpl] and [reflexivity] don't do the job, because the proofs aren't just a matter of simplifying computations. *) Theorem test_le1 : 3 <= 3. Proof. (* WORK IN CLASS *) Admitted. Theorem test_le2 : 3 <= 6. Proof. (* WORK IN CLASS *) Admitted. Theorem test_le3 : ~ (2 <= 1). Proof. (* WORK IN CLASS *) Admitted. (** The "strictly less than" relation [n < m] can now be defined in terms of [le]. *) Definition lt (n m:nat) := le (S n) m. Notation "m < n" := (lt m n). (** A few more simple relations on numbers *) Inductive square_of : nat -> nat -> Prop := sq : forall n:nat, square_of n (mult n n). Inductive next_nat (n:nat) : nat -> Prop := | nn : next_nat n (S n). Inductive next_even (n:nat) : nat -> Prop := | ne_1 : ev (S n) -> next_even n (S n) | ne_2 : ev (S (S n)) -> next_even n (S (S n)). (** **** Exercise: 2 stars *) (** Define an inductive relation [total_relation] that holds betfween every pair of natural numbers. *) (* FILL IN HERE *) (** [] *) (** **** Exercise: 2 stars *) (** Define an inductive relation [empty_relation] (on numbers) that never holds. *) (* FILL IN HERE *) (** [] *) (** **** Exercise: 3 stars (R_provability) *) Module R. (** We can define three-place relations, four-place relations, etc., in just the same way as binary relations. For example, consider the following three-place relation on numbers: *) Inductive R : nat -> nat -> nat -> Prop := | c1 : R 0 0 0 | c2 : forall m n o, R m n o -> R (S m) n (S o) | c3 : forall m n o, R m n o -> R m (S n) (S o) | c4 : forall m n o, R (S m) (S n) (S (S o)) -> R m n o | c5 : forall m n o, R m n o -> R n m o. (** - Which of the following propositions are provable? - [R 1 1 2] - [R 2 2 6] - If we dropped constructor [c5] from the definition of [R], would the set of provable propositions change? Briefly (1 sentence) explain your answer. - If we dropped constructor [c4] from the definition of [R], would the set of provable propositions change? Briefly (1 sentence) explain your answer. (* FILL IN HERE *) [] *) End R. (** **** Exercise: 4 stars, challenge problem (filter_challenge) *) (** One of the main purposes of Coq is to prove that programs match their specifications. To this end, let's prove that our definition of [filter] matches a specification. Here is the specification, written out informally in English. Suppose we have a set [X], a function [test: X->bool], and a list [l] of type [list X]. Suppose further that [l] is an in-order merge of two lists, [l1] and [l2], such that every item in [l1] satisfies [test] and no item in [l2] satisfies test. Then [filter test l = l1]. Your job is to translate this specification into a Coq theorem and prove it. (Hint: You'll need to begin by defining what it means for one list to be a merge of two others. Do this with an inductive relation, not a [Fixpoint].) *) (* FILL IN HERE *) (** [] *) (* ######################################################### *) (** * Relations, in General *) (** We've now defined a few particular relations. As you probably remember from your undergraduate discrete math course, there are a lot of ways of discussing and describing relations _in general_ -- ways of classifying relations (are they reflexive, transitive, etc.), theorems that can be proved generically about classes of relations, constructions that build one relation from another, etc. Let us pause here to review a few of these that will be useful in what follows. *) (** A _relation_ on a set [X] is a proposition parameterized by two [X]s -- i.e., it is a logical assertion involving two values from the set [X]. (This definition is just giving a name to the idea that we've been using for the past few minutes.) *) Definition relation (X: Type) := X->X->Prop. (** A relation [R] on a set [X] is a _partial function_ if, for every [x], there is at most one [y] such that [R x y] -- or, to put it differently, if [R x y1] and [R x y2] together imply [y1 = y2]. *) Definition partial_function (X: Type) (R: relation X) := forall x y1 y2 : X, R x y1 -> R x y2 -> y1 = y2. Implicit Arguments partial_function [X]. (* Check partial_function. *) (** Now, we can prove that the next_nat relation is a partial function... *) Theorem next_nat_partial_function : partial_function next_nat. Proof. unfold partial_function. intros x y1 y2 P Q. inversion P. inversion Q. reflexivity. Qed. (** ...but [<=] is not a partial function. *) Theorem le_not_a_partial_function : ~ (partial_function le). Proof. (* WORK IN CLASS *) Admitted. (** **** Exercise: 2 stars *) (** Show that the [total_relation] you defined above is not a partial function. *) (* FILL IN HERE *) (** [] *) (** **** Exercise: 2 stars *) (** Show that the [empty_relation] you defined above is a partial function. *) (* FILL IN HERE *) (** [] *) Definition reflexive (X: Type) (R: relation X) := forall a : X, R a a. Implicit Arguments reflexive [X]. Theorem le_reflexive : reflexive le. Proof. unfold reflexive. intros n. apply le_n. Qed. Definition transitive (X: Type) (R: relation X) := forall a b c : X, (R a b) -> (R b c) -> (R a c). Implicit Arguments transitive [X]. Theorem le_trans : transitive le. Proof. intros n m o Hnm Hmo. induction Hmo. Case "le_n". apply Hnm. Case "le_S". apply le_S. apply IHHmo. Qed. Theorem lt_trans: transitive lt. Proof. unfold lt. unfold transitive. intros n m o Hnm Hmo. apply le_S in Hnm. apply le_trans with (a := (S n)) (b := (S m)) (c := o). apply Hnm. apply Hmo. Qed. (** **** Exercise: 2 stars *) (** We can also prove lt_trans more laboriously by induction, without using le_trans. Do this.*) Theorem lt_trans' : transitive lt. Proof. (* Prove this by induction on evidence that [m] is lesfs than [o] *) unfold lt. unfold transitive. intros n m o Hnm Hmo. induction Hmo as [| m' Hm'o]. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 2 stars *) (** Prove the same thing again by induction on [o] *) Theorem lt_trans'' : transitive lt. Proof. unfold lt. unfold transitive. intros n m o Hnm Hmo. induction o as [| o']. (* FILL IN HERE *) Admitted. (** [] *) (** The transitivity of [le], in turn, can be used to prove some facts that will be useful later (e.g., for the proof of antisymmetry in the next section)... *) Theorem le_Sn_le : forall n m, S n <= m -> n <= m. Proof. intros n m H. apply le_trans with (S n). apply le_S. apply le_n. apply H. Qed. (** **** Exercise: 1 star *) Theorem le_S_n : forall n m, (S n <= S m) -> (n <= m). Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 2 stars (le_Sn_n_inf) *) (** Provide an informal proof of the following theorem: Theorem: For every [n], [~(S n <= n)] A formal proof of this is an optional exercise below, but try the informal proof without doing the formal proof first Proof: (* FILL IN HERE *) [] *) (** **** Exercise: 1 star *) Theorem le_Sn_n : forall n, ~ (S n <= n). Proof. (* FILL IN HERE *) Admitted. (** [] *) Definition symmetric (X: Type) (R: relation X) := forall a b : X, (R a b) -> (R b a). Implicit Arguments symmetric [X]. (** **** Exercise: 2 stars *) Theorem le_not_symmetric : ~ (symmetric le). Proof. (* FILL IN HERE *) Admitted. (** [] *) Definition antisymmetric (X: Type) (R: relation X) := forall a b : X, (R a b) -> (R b a) -> a = b. Implicit Arguments antisymmetric [X]. (** **** Exercise: 2 stars *) Theorem le_antisymmetric : antisymmetric le. Proof. (* FILL IN HERE *) Admitted. (** [] *) Definition equivalence (X:Type) (R: relation X) := (reflexive R) /\ (symmetric R) /\ (transitive R). Implicit Arguments equivalence [X]. (** A relation is called a "partial order" when it's reflexive, antisymmetric, and transitive. In the Coq standard library it's just "order" for short. *) Definition order (X:Type) (R: relation X) := (reflexive R) /\ (antisymmetric R) /\ (transitive R). Implicit Arguments order [X]. (** A preorder is almost like a partial order, but doesn't have to be antisymmetric. *) Definition preorder (X:Type) (R: relation X) := (reflexive R) /\ (transitive R). Implicit Arguments preorder [X]. Theorem le_order : order le. Proof. unfold order. split. Case "refl". apply le_reflexive. split. Case "antisym". apply le_antisymmetric. Case "transitive.". apply le_trans. Qed. (* ######################################################### *) (** * More facts about [<=] and [<] *) (** Let's pause briefly to record several facts about the [<=] and [<] relations that we are going to need later in the course. The proofs make good practice exercises. *) (** **** Exercise: 2 stars, optional (le_exercises) *) Theorem O_le_n : forall n, 0 <= n. Proof. (* FILL IN HERE *) Admitted. Theorem n_le_m__Sn_le_Sm : forall n m, n <= m -> S n <= S m. Proof. (* FILL IN HERE *) Admitted. Theorem le_plus_l : forall a b, a <= a + b. Proof. (* FILL IN HERE *) Admitted. Theorem plus_lt : forall n1 n2 m, plus n1 n2 < m -> n1 < m /\ n2 < m. Proof. (* FILL IN HERE *) Admitted. Theorem lt_S : forall n m, n < m -> n < S m. Proof. (* FILL IN HERE *) Admitted. Theorem le_step : forall n m p, n < m -> m <= S p -> n <= p. Proof. (* FILL IN HERE *) Admitted. Theorem ble_nat_true : forall n m, ble_nat n m = true -> n <= m. Proof. (* FILL IN HERE *) Admitted. Theorem ble_nat_n_Sn_false : forall n m, ble_nat n (S m) = false -> ble_nat n m = false. Proof. (* FILL IN HERE *) Admitted. (** Hint: Do the right induction! *) Theorem ble_nat_false : forall n m, ble_nat n m = false -> ~(n <= m). Proof. (* FILL IN HERE *) Admitted. (** [] *)