(** * Logic: Logic in Coq *) Set Warnings "-notation-overridden,-parsing". From LF Require Export Tactics. (** We have seen many examples of factual claims (_propositions_) and ways of presenting evidence of their truth (_proofs_). In particular, we have worked extensively with _equality propositions_ ([e1 = e2]), implications ([P -> Q]), and quantified propositions ([forall x, P]). In this chapter, we will see how Coq can be used to carry out other familiar forms of logical reasoning. Before diving into details, let's talk a bit about the status of mathematical statements in Coq. Recall that Coq is a _typed_ language, which means that every sensible expression in its world has an associated type. Logical claims are no exception: any statement we might try to prove in Coq has a type, namely [Prop], the type of _propositions_. We can see this with the [Check] command: *) Check 3 = 3 : Prop. Check forall n m : nat, n + m = m + n : Prop. (** Note that _all_ syntactically well-formed propositions have type [Prop] in Coq, regardless of whether they are true. *) (** Simply _being_ a proposition is one thing; being _provable_ is something else! *) Check 2 = 2 : Prop. Check 3 = 2 : Prop. Check forall n : nat, n = 2 : Prop. (** Indeed, propositions not only have types: they are _first-class_ entities that can be manipulated in all the same ways as any of the other things in Coq's world. *) (** So far, we've seen one primary place that propositions can appear: in [Theorem] (and [Lemma] and [Example]) declarations. *) Theorem plus_2_2_is_4 : 2 + 2 = 4. Proof. reflexivity. Qed. (** But propositions can be used in many other ways. For example, we can give a name to a proposition using a [Definition], just as we have given names to other kinds of expressions. *) Definition plus_claim : Prop := 2 + 2 = 4. Check plus_claim : Prop. (** We can later use this name in any situation where a proposition is expected -- for example, as the claim in a [Theorem] declaration. *) Theorem plus_claim_is_true : plus_claim. Proof. reflexivity. Qed. (** We can also write _parameterized_ propositions -- that is, functions that take arguments of some type and return a proposition. *) (** For instance, the following function takes a number and returns a proposition asserting that this number is equal to three: *) Definition is_three (n : nat) : Prop := n = 3. Check is_three : nat -> Prop. (** In Coq, functions that return propositions are said to define _properties_ of their arguments. For instance, here's a (polymorphic) property defining the familiar notion of an _injective function_. *) Definition injective {A B} (f : A -> B) := forall x y : A, f x = f y -> x = y. Lemma succ_inj : injective S. Proof. intros n m H. injection H as H1. apply H1. Qed. (** The equality operator [=] is also a function that returns a [Prop]. The expression [n = m] is syntactic sugar for [eq n m] (defined in Coq's standard library using the [Notation] mechanism). Because [eq] can be used with elements of any type, it is also polymorphic: *) Check @eq : forall A : Type, A -> A -> Prop. (** (Notice that we wrote [@eq] instead of [eq]: The type argument [A] to [eq] is declared as implicit, and we need to turn off the inference of this implicit argument to see the full type of [eq].) *) (* ################################################################# *) (** * Logical Connectives *) (* ================================================================= *) (** ** Conjunction *) (** The _conjunction_, or _logical and_, of propositions [A] and [B] is written [A /\ B], representing the claim that both [A] and [B] are true. *) Example and_example : 3 + 4 = 7 /\ 2 * 2 = 4. (** To prove a conjunction, use the [split] tactic. It will generate two subgoals, one for each part of the statement: *) Proof. split. - (* 3 + 4 = 7 *) reflexivity. - (* 2 * 2 = 4 *) reflexivity. Qed. (** For any propositions [A] and [B], if we assume that [A] is true and that [B] is true, we can conclude that [A /\ B] is also true. *) Lemma and_intro : forall A B : Prop, A -> B -> A /\ B. Proof. intros A B HA HB. split. - apply HA. - apply HB. Qed. (** Since applying a theorem with hypotheses to some goal has the effect of generating as many subgoals as there are hypotheses for that theorem, we can apply [and_intro] to achieve the same effect as [split]. *) Example and_example' : 3 + 4 = 7 /\ 2 * 2 = 4. Proof. apply and_intro. - (* 3 + 4 = 7 *) reflexivity. - (* 2 + 2 = 4 *) reflexivity. Qed. (** **** Exercise: 2 stars, standard, optional (and_exercise) *) Example and_exercise : forall n m : nat, n + m = 0 -> n = 0 /\ m = 0. Proof. (* FILL IN HERE *) Admitted. (** [] *) (** So much for proving conjunctive statements. To go in the other direction -- i.e., to _use_ a conjunctive hypothesis to help prove something else -- we employ the [destruct] tactic. If the proof context contains a hypothesis [H] of the form [A /\ B], writing [destruct H as [HA HB]] will remove [H] from the context and add two new hypotheses: [HA], stating that [A] is true, and [HB], stating that [B] is true. *) Lemma and_example2 : forall n m : nat, n = 0 /\ m = 0 -> n + m = 0. Proof. (* WORKED IN CLASS *) intros n m H. destruct H as [Hn Hm]. rewrite Hn. rewrite Hm. reflexivity. Qed. (** As usual, we can also destruct [H] right when we introduce it, instead of introducing and then destructing it: *) Lemma and_example2' : forall n m : nat, n = 0 /\ m = 0 -> n + m = 0. Proof. intros n m [Hn Hm]. rewrite Hn. rewrite Hm. reflexivity. Qed. (** You may wonder why we bothered packing the two hypotheses [n = 0] and [m = 0] into a single conjunction, since we could have also stated the theorem with two separate premises: *) Lemma and_example2'' : forall n m : nat, n = 0 -> m = 0 -> n + m = 0. Proof. intros n m Hn Hm. rewrite Hn. rewrite Hm. reflexivity. Qed. (** For this specific theorem, both formulations are fine. But it's important to understand how to work with conjunctive hypotheses because conjunctions often arise from intermediate steps in proofs, especially in larger developments. Here's a simple example: *) Lemma and_example3 : forall n m : nat, n + m = 0 -> n * m = 0. Proof. (* WORKED IN CLASS *) intros n m H. apply and_exercise in H. destruct H as [Hn Hm]. rewrite Hn. reflexivity. Qed. (** Another common situation with conjunctions is that we know [A /\ B] but in some context we need just [A] or just [B]. In such cases we can do a [destruct] (possibly as part of an [intros]) and use an underscore pattern [_] to indicate that the unneeded conjunct should just be thrown away. *) Lemma proj1 : forall P Q : Prop, P /\ Q -> P. Proof. intros P Q HPQ. destruct HPQ as [HP _]. apply HP. Qed. (** **** Exercise: 1 star, standard, optional (proj2) *) Lemma proj2 : forall P Q : Prop, P /\ Q -> Q. Proof. (* FILL IN HERE *) Admitted. (** [] *) (** Finally, we sometimes need to rearrange the order of conjunctions and/or the grouping of multi-way conjunctions. The following commutativity and associativity theorems are handy in such cases. *) Theorem and_commut : forall P Q : Prop, P /\ Q -> Q /\ P. Proof. intros P Q [HP HQ]. split. - (* left *) apply HQ. - (* right *) apply HP. Qed. (** **** Exercise: 2 stars, standard, optional (and_assoc) (In the following proof of associativity, notice how the _nested_ [intros] pattern breaks the hypothesis [H : P /\ (Q /\ R)] down into [HP : P], [HQ : Q], and [HR : R]. Finish the proof from there.) *) Theorem and_assoc : forall P Q R : Prop, P /\ (Q /\ R) -> (P /\ Q) /\ R. Proof. intros P Q R [HP [HQ HR]]. (* FILL IN HERE *) Admitted. (** [] *) (** By the way, the infix notation [/\] is actually just syntactic sugar for [and A B]. That is, [and] is a Coq operator that takes two propositions as arguments and yields a proposition. *) Check and : Prop -> Prop -> Prop. (* ================================================================= *) (** ** Disjunction *) (** Another important connective is the _disjunction_, or _logical or_, of two propositions: [A \/ B] is true when either [A] or [B] is. (This infix notation stands for [or A B], where [or : Prop -> Prop -> Prop].) *) (** To use a disjunctive hypothesis in a proof, we proceed by case analysis (which, as with other data types like [nat], can be done explicitly with [destruct] or implicitly with an [intros] pattern): *) Lemma eq_mult_0 : forall n m : nat, n = 0 \/ m = 0 -> n * m = 0. Proof. (* This pattern implicitly does case analysis on [n = 0 \/ m = 0] *) intros n m [Hn | Hm]. - (* Here, [n = 0] *) rewrite Hn. reflexivity. - (* Here, [m = 0] *) rewrite Hm. rewrite <- mult_n_O. reflexivity. Qed. (** Conversely, to show that a disjunction holds, it suffices to show that one of its sides holds. This is done via two tactics, [left] and [right]. As their names imply, the first one requires proving the left side of the disjunction, while the second requires proving its right side. Here is a trivial use... *) Lemma or_intro_l : forall A B : Prop, A -> A \/ B. Proof. intros A B HA. left. apply HA. Qed. (** ... and here is a slightly more interesting example requiring both [left] and [right]: *) Lemma zero_or_succ : forall n : nat, n = 0 \/ n = S (pred n). Proof. (* WORKED IN CLASS *) intros [|n']. - left. reflexivity. - right. reflexivity. Qed. (** **** Exercise: 1 star, standard (mult_eq_0) *) Lemma mult_eq_0 : forall n m, n * m = 0 -> n = 0 \/ m = 0. Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 1 star, standard (or_commut) *) Theorem or_commut : forall P Q : Prop, P \/ Q -> Q \/ P. Proof. (* FILL IN HERE *) Admitted. (** [] *) (* ================================================================= *) (** ** Falsehood and Negation So far, we have mostly been concerned with proving that certain things are _true_ -- addition is commutative, appending lists is associative, etc. Of course, we may also be interested in negative results, demonstrating that some given proposition is _not_ true. Such statements are expressed with the logical negation operator [~]. *) (** To see how negation works, recall the _principle of explosion_ from the [Tactics] chapter, which asserts that, if we assume a contradiction, then any other proposition can be derived. Following this intuition, we could define [~ P] ("not [P]") as [forall Q, P -> Q]. Coq actually makes a slightly different (but equivalent) choice, defining [~ P] as [P -> False], where [False] is a specific contradictory proposition defined in the standard library. *) Module MyNot. Definition not (P:Prop) := P -> False. Notation "~ x" := (not x) : type_scope. Check not : Prop -> Prop. End MyNot. (** Since [False] is a contradictory proposition, the principle of explosion also applies to it. If we get [False] into the proof context, we can use [destruct] on it to complete any goal: *) Theorem ex_falso_quodlibet : forall (P:Prop), False -> P. Proof. (* WORKED IN CLASS *) intros P contra. destruct contra. Qed. (** The Latin _ex falso quodlibet_ means, literally, "from falsehood follows whatever you like"; this is another common name for the principle of explosion. *) (** **** Exercise: 2 stars, standard, optional (not_implies_our_not) Show that Coq's definition of negation implies the intuitive one mentioned above: *) Fact not_implies_our_not : forall (P:Prop), ~ P -> (forall (Q:Prop), P -> Q). Proof. (* FILL IN HERE *) Admitted. (** [] *) (** Inequality is a frequent enough example of negated statement that there is a special notation for it, [x <> y]: Notation "x <> y" := (~(x = y)). *) (** We can use [not] to state that [0] and [1] are different elements of [nat]: *) Theorem zero_not_one : 0 <> 1. Proof. (** The proposition [0 <> 1] is exactly the same as [~(0 = 1)], that is [not (0 = 1)], which unfolds to [(0 = 1) -> False]. (We use [unfold not] explicitly here to illustrate that point, but generally it can be omitted.) *) unfold not. (** To prove an inequality, we may assume the opposite equality... *) intros contra. (** ... and deduce a contradiction from it. Here, the equality [O = S O] contradicts the disjointness of constructors [O] and [S], so [discriminate] takes care of it. *) discriminate contra. Qed. (** It takes a little practice to get used to working with negation in Coq. Even though you can see perfectly well why a statement involving negation is true, it can be a little tricky at first to make Coq understand it! Here are proofs of a few familiar facts to get you warmed up. *) Theorem not_False : ~ False. Proof. unfold not. intros H. destruct H. Qed. Theorem contradiction_implies_anything : forall P Q : Prop, (P /\ ~P) -> Q. Proof. (* WORKED IN CLASS *) intros P Q [HP HNA]. unfold not in HNA. apply HNA in HP. destruct HP. Qed. Theorem double_neg : forall P : Prop, P -> ~~P. Proof. (* WORKED IN CLASS *) intros P H. unfold not. intros G. apply G. apply H. Qed. (** **** Exercise: 2 stars, advanced, optional (double_neg_inf) Write an informal proof of [double_neg]: _Theorem_: [P] implies [~~P], for any proposition [P]. *) (* FILL IN HERE *) (* Do not modify the following line: *) Definition manual_grade_for_double_neg_inf : option (nat*string) := None. (** [] *) (** **** Exercise: 2 stars, standard, especially useful (contrapositive) *) Theorem contrapositive : forall (P Q : Prop), (P -> Q) -> (~Q -> ~P). Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 1 star, standard, optional (not_both_true_and_false) *) Theorem not_both_true_and_false : forall P : Prop, ~ (P /\ ~P). Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 1 star, advanced, optional (informal_not_PNP) Write an informal proof (in English) of the proposition [forall P : Prop, ~(P /\ ~P)]. *) (* FILL IN HERE *) (* Do not modify the following line: *) Definition manual_grade_for_informal_not_PNP : option (nat*string) := None. (** [] *) (** Since inequality involves a negation, it also requires a little practice to be able to work with it fluently. Here is one useful trick. If you are trying to prove a goal that is nonsensical (e.g., the goal state is [false = true]), apply [ex_falso_quodlibet] to change the goal to [False]. This makes it easier to use assumptions of the form [~P] that may be available in the context -- in particular, assumptions of the form [x<>y]. *) Theorem not_true_is_false : forall b : bool, b <> true -> b = false. Proof. intros b H. destruct b eqn:HE. - (* b = true *) unfold not in H. apply ex_falso_quodlibet. apply H. reflexivity. - (* b = false *) reflexivity. Qed. (** Since reasoning with [ex_falso_quodlibet] is quite common, Coq provides a built-in tactic, [exfalso], for applying it. *) Theorem not_true_is_false' : forall b : bool, b <> true -> b = false. Proof. intros [] H. (* note implicit [destruct b] here *) - (* b = true *) unfold not in H. exfalso. (* <=== *) apply H. reflexivity. - (* b = false *) reflexivity. Qed. (* ================================================================= *) (** ** Truth *) (** Besides [False], Coq's standard library also defines [True], a proposition that is trivially true. To prove it, we use the predefined constant [I : True]: *) Lemma True_is_true : True. Proof. apply I. Qed. (** Unlike [False], which is used extensively, [True] is used quite rarely, since it is trivial (and therefore uninteresting) to prove as a goal, and it carries no useful information as a hypothesis. But it can be quite useful when defining complex [Prop]s using conditionals or as a parameter to higher-order [Prop]s. We will see examples later on. *) (* ================================================================= *) (** ** Logical Equivalence *) (** The handy "if and only if" connective, which asserts that two propositions have the same truth value, is simply the conjunction of two implications. *) Module MyIff. Definition iff (P Q : Prop) := (P -> Q) /\ (Q -> P). Notation "P <-> Q" := (iff P Q) (at level 95, no associativity) : type_scope. End MyIff. Theorem iff_sym : forall P Q : Prop, (P <-> Q) -> (Q <-> P). Proof. (* WORKED IN CLASS *) intros P Q [HAB HBA]. split. - (* -> *) apply HBA. - (* <- *) apply HAB. Qed. Lemma not_true_iff_false : forall b, b <> true <-> b = false. Proof. (* WORKED IN CLASS *) intros b. split. - (* -> *) apply not_true_is_false. - (* <- *) intros H. rewrite H. intros H'. discriminate H'. Qed. (** **** Exercise: 1 star, standard, optional (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 P : Prop, P <-> P. Proof. (* FILL IN HERE *) Admitted. Theorem iff_trans : forall P Q R : Prop, (P <-> Q) -> (Q <-> R) -> (P <-> R). Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 3 stars, standard (or_distributes_over_and) *) Theorem or_distributes_over_and : forall P Q R : Prop, P \/ (Q /\ R) <-> (P \/ Q) /\ (P \/ R). Proof. (* FILL IN HERE *) Admitted. (** [] *) (* ================================================================= *) (** ** Setoids and Logical Equivalence *) (** Some of Coq's tactics treat [iff] statements specially, avoiding the need for some low-level proof-state manipulation. In particular, [rewrite] and [reflexivity] can be used with [iff] statements, not just equalities. To enable this behavior, we have to import the Coq library that supports it: *) From Coq Require Import Setoids.Setoid. (** A "setoid" is a set equipped with an equivalence relation, that is, a relation that is reflexive, symmetric, and transitive. When two elements of a set are equivalent according to the relation, [rewrite] can be used to replace one element with the other. We've seen that already with the equality relation [=] in Coq: when [x = y], we can use [rewrite] to replace [x] with [y], or vice-versa. Similarly, the logical equivalence relation [<->] is reflexive, symmetric, and transitive, so we can use it to replace one part of a proposition with another: if [P <-> Q], then we can use [rewrite] to replace [P] with [Q], or vice-versa. *) (** Here is a simple example demonstrating how these tactics work with [iff]. First, let's prove a couple of basic iff equivalences. *) Lemma mult_0 : forall n m, n * m = 0 <-> n = 0 \/ m = 0. Proof. split. - apply mult_eq_0. - apply eq_mult_0. Qed. Theorem or_assoc : forall P Q R : Prop, P \/ (Q \/ R) <-> (P \/ Q) \/ R. Proof. intros P Q R. split. - intros [H | [H | H]]. + left. left. apply H. + left. right. apply H. + right. apply H. - intros [[H | H] | H]. + left. apply H. + right. left. apply H. + right. right. apply H. Qed. (** We can now use these facts with [rewrite] and [reflexivity] to give smooth proofs of statements involving equivalences. For example, here is a ternary version of the previous [mult_0] result: *) Lemma mult_0_3 : forall n m p, n * m * p = 0 <-> n = 0 \/ m = 0 \/ p = 0. Proof. intros n m p. rewrite mult_0. rewrite mult_0. rewrite or_assoc. reflexivity. Qed. (** The [apply] tactic can also be used with [<->]. When given an equivalence as its argument, [apply] tries to guess which direction of the equivalence will be useful. *) Lemma apply_iff_example : forall n m : nat, n * m = 0 -> n = 0 \/ m = 0. Proof. intros n m H. apply mult_0. apply H. Qed. (* ================================================================= *) (** ** Existential Quantification *) (** Another important logical connective is _existential quantification_. To say that there is some [x] of type [T] such that some property [P] holds of [x], we write [exists x : T, P]. As with [forall], the type annotation [: T] can be omitted if Coq is able to infer from the context what the type of [x] should be. *) (** To prove a statement of the form [exists x, P], we must show that [P] holds for some specific choice of value for [x], known as the _witness_ of the existential. This is done in two steps: First, we explicitly tell Coq which witness [t] we have in mind by invoking the tactic [exists t]. Then we prove that [P] holds after all occurrences of [x] are replaced by [t]. *) Definition even x := exists n : nat, x = double n. Lemma four_is_even : even 4. Proof. unfold even. exists 2. reflexivity. Qed. (** Conversely, if we have an existential hypothesis [exists x, P] in the context, we can destruct it to obtain a witness [x] and a hypothesis stating that [P] holds of [x]. *) Theorem exists_example_2 : forall n, (exists m, n = 4 + m) -> (exists o, n = 2 + o). Proof. (* WORKED IN CLASS *) intros n [m Hm]. (* note implicit [destruct] here *) exists (2 + m). apply Hm. Qed. (** **** Exercise: 1 star, standard, especially useful (dist_not_exists) Prove that "[P] holds for all [x]" implies "there is no [x] for which [P] does not hold." (Hint: [destruct H as [x E]] works on existential assumptions!) *) 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, standard (dist_exists_or) 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. (** [] *) (* ################################################################# *) (** * Programming with Propositions *) (** The logical connectives that we have seen provide a rich vocabulary for defining complex propositions from simpler ones. To illustrate, let's look at how to express the claim that an element [x] occurs in a list [l]. Notice that this property has a simple recursive structure: - If [l] is the empty list, then [x] cannot occur in it, so the property "[x] appears in [l]" is simply false. - Otherwise, [l] has the form [x' :: l']. In this case, [x] occurs in [l] if either it is equal to [x'] or it occurs in [l']. *) (** We can translate this directly into a straightforward recursive function taking an element and a list and returning a proposition (!): *) Fixpoint In {A : Type} (x : A) (l : list A) : Prop := match l with | [] => False | x' :: l' => x' = x \/ In x l' end. (** When [In] is applied to a concrete list, it expands into a concrete sequence of nested disjunctions. *) Example In_example_1 : In 4 [1; 2; 3; 4; 5]. Proof. (* WORKED IN CLASS *) simpl. right. right. right. left. reflexivity. Qed. Example In_example_2 : forall n, In n [2; 4] -> exists n', n = 2 * n'. Proof. (* WORKED IN CLASS *) simpl. intros n [H | [H | []]]. - exists 1. rewrite <- H. reflexivity. - exists 2. rewrite <- H. reflexivity. Qed. (** (Notice the use of the empty pattern to discharge the last case _en passant_.) *) (** We can also prove more generic, higher-level lemmas about [In]. Note, in the first, how [In] starts out applied to a variable and only gets expanded when we do case analysis on this variable: *) Theorem In_map : forall (A B : Type) (f : A -> B) (l : list A) (x : A), In x l -> In (f x) (map f l). Proof. intros A B f l x. induction l as [|x' l' IHl']. - (* l = nil, contradiction *) simpl. intros []. - (* l = x' :: l' *) simpl. intros [H | H]. + rewrite H. left. reflexivity. + right. apply IHl'. apply H. Qed. (** This way of defining propositions recursively, though convenient in some cases, also has some drawbacks. In particular, it is subject to Coq's usual restrictions regarding the definition of recursive functions, e.g., the requirement that they be "obviously terminating." In the next chapter, we will see how to define propositions _inductively_, a different technique with its own set of strengths and limitations. *) (** **** Exercise: 3 stars, standard (In_map_iff) *) Theorem In_map_iff : forall (A B : Type) (f : A -> B) (l : list A) (y : B), In y (map f l) <-> exists x, f x = y /\ In x l. Proof. intros A B f l y. split. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 2 stars, standard (In_app_iff) *) Theorem In_app_iff : forall A l l' (a:A), In a (l++l') <-> In a l \/ In a l'. Proof. intros A l. induction l as [|a' l' IH]. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 3 stars, standard, especially useful (All) Recall that functions returning propositions can be seen as _properties_ of their arguments. For instance, if [P] has type [nat -> Prop], then [P n] states that property [P] holds of [n]. Drawing inspiration from [In], write a recursive function [All] stating that some property [P] holds of all elements of a list [l]. To make sure your definition is correct, prove the [All_In] lemma below. (Of course, your definition should _not_ just restate the left-hand side of [All_In].) *) Fixpoint All {T : Type} (P : T -> Prop) (l : list T) : Prop (* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted. Theorem All_In : forall T (P : T -> Prop) (l : list T), (forall x, In x l -> P x) <-> All P l. Proof. (* FILL IN HERE *) Admitted. (** [] *) (** **** Exercise: 2 stars, standard, optional (combine_odd_even) Complete the definition of the [combine_odd_even] function below. It takes as arguments two properties of numbers, [Podd] and [Peven], and it should return a property [P] such that [P n] is equivalent to [Podd n] when [n] is odd and equivalent to [Peven n] otherwise. *) Definition combine_odd_even (Podd Peven : nat -> Prop) : nat -> Prop (* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted. (** To test your definition, prove the following facts: *) Theorem combine_odd_even_intro : forall (Podd Peven : nat -> Prop) (n : nat), (oddb n = true -> Podd n) -> (oddb n = false -> Peven n) -> combine_odd_even Podd Peven n. Proof. (* FILL IN HERE *) Admitted. Theorem combine_odd_even_elim_odd : forall (Podd Peven : nat -> Prop) (n : nat), combine_odd_even Podd Peven n -> oddb n = true -> Podd n. Proof. (* FILL IN HERE *) Admitted. Theorem combine_odd_even_elim_even : forall (Podd Peven : nat -> Prop) (n : nat), combine_odd_even Podd Peven n -> oddb n = false -> Peven n. Proof. (* FILL IN HERE *) Admitted. (** [] *)