(* midterm_project.v *) (* LPP 2024 - CS3234 2023-2024, Sem2 *) (* Olivier Danvy *) (* Version of 22 Feb 2024 *) (* ********** *) (* A study of polymorphic lists. *) (* members of the group: Kyriel Mortel Abad (A0258247W, kyriel@u.nus.edu) Kyle Zheng Ching Chan (A0174144J, e0205139@u.nus.edu) Yadunand Prem (A0253252M, y.p@u.nus.edu) Wong Kok Rui (A0174107L, kokrui@u.nus.edu) Koh Chan Hong (A0254616A, e0959286@u.nus.edu) Fong Yi Yong Calvin (A0255291A, ccc@u.nus.edu) Chandrasekaran Akash (A0152223W, c.akash@u.nus.edu) date: 21 March 2024 please upload one .v file and one .pdf file containing a project report desiderata: - the file should be readable, i.e., properly indented and using items or {...} for subgoals - each use of a tactic should achieve one proof step - all lemmas should be applied to all their arguments - there should be no occurrences of admit, Admitted, and Abort *) (* ********** *) (* Paraphernalia: *) Ltac fold_unfold_tactic name := intros; unfold name; fold name; reflexivity. Require Import Arith Bool List. (* ********** *) Definition make_singleton_list (V : Type) (v : V) : list V := v :: nil. (* ********** *) Definition eqb_option (V : Type) (eqb_V : V -> V -> bool) (ov1 ov2 : option V) : bool := match ov1 with Some v1 => match ov2 with Some v2 => eqb_V v1 v2 | None => false end | None => match ov2 with Some v2 => false | None => true end end. (* ********** *) Fixpoint eqb_list (V : Type) (eqb_V : V -> V -> bool) (v1s v2s : list V) : bool := match v1s with nil => match v2s with nil => true | v2 :: v2s' => false end | v1 :: v1s' => match v2s with nil => false | v2 :: v2s' => eqb_V v1 v2 && eqb_list V eqb_V v1s' v2s' end end. Lemma fold_unfold_eqb_list_nil : forall (V : Type) (eqb_V : V -> V -> bool) (v2s : list V), eqb_list V eqb_V nil v2s = match v2s with nil => true | v2 :: v2s' => false end. Proof. fold_unfold_tactic eqb_list. Qed. Lemma fold_unfold_eqb_list_cons : forall (V : Type) (eqb_V : V -> V -> bool) (v1 : V) (v1s' v2s : list V), eqb_list V eqb_V (v1 :: v1s') v2s = match v2s with nil => false | v2 :: v2s' => eqb_V v1 v2 && eqb_list V eqb_V v1s' v2s' end. Proof. fold_unfold_tactic eqb_list. Qed. (* ***** *) (* Task 1: *) (* {task_1} *) Theorem soundness_of_equality_over_lists : forall (V : Type) (eqb_V : V -> V -> bool), (forall v1 v2 : V, eqb_V v1 v2 = true -> v1 = v2) -> forall v1s v2s : list V, eqb_list V eqb_V v1s v2s = true -> v1s = v2s. Proof. intros V eqb_V S_eqb_V v1s. induction v1s as [ | v1 v1s' IHv1s']. - intros [ | v2 v2s']. + reflexivity. + rewrite -> (fold_unfold_eqb_list_nil V eqb_V (v2 :: v2s')). intro H_absurd. discriminate H_absurd. - intros [ | v2 v2s']. + rewrite -> (fold_unfold_eqb_list_cons V eqb_V v1 v1s' nil). intro H_absurd. discriminate H_absurd. + rewrite -> (fold_unfold_eqb_list_cons V eqb_V v1 v1s' (v2 :: v2s')). intro eqb_v1_v2_and_eqb_v1s'_v2s'. destruct (andb_true_iff (eqb_V v1 v2) (eqb_list V eqb_V v1s' v2s')) as [H_tmp _]. destruct (H_tmp eqb_v1_v2_and_eqb_v1s'_v2s') as [eqb_v1_v2 eqb_v1s'_v2s']. rewrite -> (IHv1s' v2s' eqb_v1s'_v2s'). rewrite -> (S_eqb_V v1 v2 eqb_v1_v2). reflexivity. Qed. Theorem completeness_of_equality_over_lists : forall (V : Type) (eqb_V : V -> V -> bool), (forall v1 v2 : V, v1 = v2 -> eqb_V v1 v2 = true) -> forall v1s v2s : list V, v1s = v2s -> eqb_list V eqb_V v1s v2s = true. Proof. intros V eqb_V C_eqb_V v1s. induction v1s as [ | v1 v1s' IHv1s']. - intros [ | v2 v2s']. + rewrite -> (fold_unfold_eqb_list_nil V eqb_V nil). reflexivity. + intro H_absurd. discriminate H_absurd. - intros [ | v2 v2s']. + intro H_absurd. discriminate H_absurd. + intro eq_v1s_v2s. rewrite -> (fold_unfold_eqb_list_cons V eqb_V v1 v1s' (v2 :: v2s')). injection eq_v1s_v2s as eq_v1_v2 eq_v1s'_v2s'. destruct (andb_true_iff (eqb_V v1 v2) (eqb_list V eqb_V v1s' v2s')) as [_ H_tmp]. apply H_tmp. split. * exact (C_eqb_V v1 v2 eq_v1_v2). * exact (IHv1s' v2s' eq_v1s'_v2s'). Qed. (* {END} *) (* ********** *) (* A study of the polymorphic length function: *) (* {specification_of_list_length} *) Definition specification_of_list_length (length : forall V : Type, list V -> nat) := (forall V : Type, length V nil = 0) /\ (forall (V : Type) (v : V) (vs' : list V), length V (v :: vs') = S (length V vs')). (* {END} *) (* Unit-test function: *) Definition test_list_length (candidate : forall V : Type, list V -> nat) := (candidate nat nil =? 0) && (candidate bool nil =? 0) && (candidate nat (1 :: nil) =? 1) && (candidate bool (true :: nil) =? 1) && (candidate nat (2 :: 1 :: nil) =? 2) && (candidate bool (false :: true :: nil) =? 2) && (candidate nat (3 :: 2 :: 1 :: nil) =? 3) && (candidate bool (false :: false :: true :: nil) =? 3). (* The specification specifies at most one length function: *) Theorem there_is_at_most_one_list_length_function : forall (V : Type) (list_length_1 list_length_2 : forall V : Type, list V -> nat), specification_of_list_length list_length_1 -> specification_of_list_length list_length_2 -> forall vs : list V, list_length_1 V vs = list_length_2 V vs. Proof. intros V list_length_1 list_length_2. unfold specification_of_list_length. intros [S_list_length_1_nil S_list_length_1_cons] [S_list_length_2_nil S_list_length_2_cons] vs. induction vs as [ | v vs' IHvs']. - Check (S_list_length_2_nil V). rewrite -> (S_list_length_2_nil V). Check (S_list_length_1_nil V). exact (S_list_length_1_nil V). - Check (S_list_length_1_cons V v vs'). rewrite -> (S_list_length_1_cons V v vs'). rewrite -> (S_list_length_2_cons V v vs'). rewrite -> IHvs'. reflexivity. Qed. (* Recursive implementation of the length function: *) Fixpoint list_length (V : Type) (vs : list V) : nat := match vs with nil => 0 | v :: vs' => S (list_length V vs') end. Compute (test_list_length list_length). (* Associated fold-unfold lemmas: *) Lemma fold_unfold_list_length_nil : forall V : Type, list_length V nil = 0. Proof. fold_unfold_tactic list_length. Qed. Lemma fold_unfold_list_length_cons : forall (V : Type) (v : V) (vs' : list V), list_length V (v :: vs') = S (list_length V vs'). Proof. fold_unfold_tactic list_length. Qed. (* The specification specifies at least one length function: *) Theorem list_length_satisfies_the_specification_of_list_length : specification_of_list_length list_length. Proof. unfold specification_of_list_length. exact (conj fold_unfold_list_length_nil fold_unfold_list_length_cons). Qed. (* ***** *) (* Task 2: *) (* {task_2} *) (* Implement the length function using an accumulator. *) Fixpoint list_length_acc (V : Type) (vs : list V) (a : nat) : nat := match vs with nil => a | v :: vs' => list_length_acc V vs' (S a) end. Definition list_length_alt (V : Type) (vs : list V) : nat := list_length_acc V vs 0. Compute (test_list_length list_length_alt). (* Associated fold-unfold lemmas: *) Lemma fold_unfold_list_length_acc_nil : forall (V : Type) (a : nat), list_length_acc V nil a = a. Proof. fold_unfold_tactic list_length_acc. Qed. Lemma fold_unfold_list_length_acc_cons : forall (V : Type) (v : V) (vs' : list V) (a : nat), list_length_acc V (v :: vs') a = list_length_acc V vs' (S a). Proof. fold_unfold_tactic list_length_acc. Qed. (* {END} *) (* {task_2b} *) Lemma about_list_length_acc : forall (V : Type) (vs : list V) (a : nat), list_length_acc V vs a = Nat.add (list_length_acc V vs 0) a. Proof. intros V vs. induction vs as [ | v vs' IHvs']. - intro a. rewrite -> (fold_unfold_list_length_acc_nil V a). rewrite -> (fold_unfold_list_length_acc_nil V 0). exact (eq_sym (Nat.add_0_l a)). - intro a. rewrite -> (fold_unfold_list_length_acc_cons V v vs' a). rewrite -> (fold_unfold_list_length_acc_cons V v vs' 0). rewrite -> (IHvs' (S a)). rewrite -> (IHvs' 1). rewrite <- (Nat.add_1_l a). exact (Nat.add_assoc (list_length_acc V vs' 0) 1 a). Qed. (* The specification specifies the alt length function too: *) Theorem list_length_alt_satisfies_the_specification_of_list_length : specification_of_list_length list_length_alt. Proof. unfold specification_of_list_length, list_length_alt. split. - intro V. exact (fold_unfold_list_length_acc_nil V 0). - intros V v vs'. rewrite -> (fold_unfold_list_length_acc_cons V v vs' 0). rewrite -> (about_list_length_acc V vs' 1). exact (Nat.add_1_r (list_length_acc V vs' 0)). Qed. (* {END} *) (* ********** *) (* A study of the polymorphic copy function: *) (* {specification_of_list_copy} *) Definition specification_of_list_copy (copy : forall V : Type, list V -> list V) := (forall V : Type, copy V nil = nil) /\ (forall (V : Type) (v : V) (vs' : list V), copy V (v :: vs') = v :: (copy V vs')). (* {END} *) Definition test_list_copy (candidate : forall V : Type, list V -> list V) := (eqb_list nat Nat.eqb (candidate nat nil) nil) && (eqb_list bool Bool.eqb (candidate bool nil) nil) && (eqb_list nat Nat.eqb (candidate nat (1 :: nil)) (1 :: nil)) && (eqb_list bool Bool.eqb (candidate bool (true :: nil)) (true :: nil)) && (eqb_list nat Nat.eqb (candidate nat (2 :: 1 :: nil)) (2 :: 1 :: nil)) && (eqb_list bool Bool.eqb (candidate bool (false :: true :: nil)) (false :: true :: nil)). (* ***** *) (* Task 3: *) (* a. Expand the unit-test function for copy with a few more tests. *) (* {task_3_a} *) Definition test_list_copy_more (candidate : forall V : Type, list V -> list V) := (eqb_list (list nat) (eqb_list nat Nat.eqb) (candidate (list nat) (nil :: nil)) (nil :: nil)) && (eqb_list (list bool) (eqb_list bool Bool.eqb) (candidate (list bool) (nil :: nil)) (nil :: nil)) && (eqb_list (list nat) (eqb_list nat Nat.eqb) (candidate (list nat) ((1 :: 2 :: nil) :: nil)) ((1 :: 2 :: nil) :: nil)) && (eqb_list (list bool) (eqb_list bool Bool.eqb) (candidate (list bool) ((true :: nil) :: nil)) ((true :: nil) :: nil)). (* {END} *) (* b. Implement the copy function recursively. *) (* {task_3_b} *) Fixpoint list_copy (V : Type) (vs : list V) : list V := match vs with nil => nil | v :: vs' => v :: (list_copy V vs') end. Compute (test_list_copy list_copy). Compute (test_list_copy_more list_copy). (* {END} *) (* c. State its associated fold-unfold lemmas. *) (* {task_3_c} *) Lemma fold_unfold_list_copy_nil : forall (V : Type), list_copy V nil = nil. Proof. fold_unfold_tactic list_copy. Qed. Lemma fold_unfold_list_copy_cons : forall (V : Type) (v : V) (vs' : list V), list_copy V (v :: vs') = v :: (list_copy V vs'). Proof. fold_unfold_tactic list_copy. Qed. (* {END} *) (* d. Prove whether your implementation satisfies the specification. *) (* {task_3_d} *) Theorem list_copy_satisfies_the_specification_of_list_copy : specification_of_list_copy list_copy. Proof. unfold specification_of_list_copy. exact (conj fold_unfold_list_copy_nil fold_unfold_list_copy_cons). Qed. (* {END} *) (* e. Prove whether copy is idempotent. *) (* {task_3_e} *) Proposition list_copy_is_idempotent : forall (V : Type) (vs : list V), list_copy V (list_copy V vs) = list_copy V vs. Proof. intros V vs. induction vs as [ | v vs' IHvs']. - rewrite -> (fold_unfold_list_copy_nil V). rewrite -> (fold_unfold_list_copy_nil V). reflexivity. - rewrite -> (fold_unfold_list_copy_cons V v vs'). rewrite -> (fold_unfold_list_copy_cons V v (list_copy V vs')). rewrite -> IHvs'. reflexivity. Qed. (* {END} *) (* f. Prove whether copying a list preserves its length. *) (* {task_3_f} *) Proposition list_copy_preserves_list_length : forall (V : Type) (vs : list V), list_length V (list_copy V vs) = list_length V vs. Proof. intros V vs. induction vs as [ | v vs' IHvs']. - rewrite -> (fold_unfold_list_copy_nil V). reflexivity. - rewrite -> (fold_unfold_list_copy_cons V v vs'). rewrite -> (fold_unfold_list_length_cons V v vs'). rewrite -> (fold_unfold_list_length_cons V v (list_copy V vs')). rewrite -> IHvs'. reflexivity. Qed. (* {END} *) (* g. Subsidiary question: can you think of a strikingly simple implementation of the copy function? if so, pray show that it satisfies the specification of copy; is it equivalent to your recursive implementation? *) (* {task_3_g} *) Definition list_copy_identity (V : Type) (vs : list V) := vs. Theorem list_copy_identity_satisfies_the_specification_of_list_copy : specification_of_list_copy list_copy_identity. Proof. unfold specification_of_list_copy, list_copy_identity. split. - intro V. reflexivity. - intros V v vs'. reflexivity. Qed. Proposition there_is_at_most_one_list_copy_function : forall (V : Type) (list_copy_1 list_copy_2 : forall V : Type, list V -> list V), specification_of_list_copy list_copy_1 -> specification_of_list_copy list_copy_2 -> forall vs : list V, list_copy_1 V vs = list_copy_2 V vs. Proof. intros V list_copy_1 list_copy_2. unfold specification_of_list_copy. intros [S_list_copy_1_nil S_list_copy_1_cons] [S_list_copy_2_nil S_list_copy_2_cons] vs. induction vs as [ | v vs' IHvs']. - rewrite -> (S_list_copy_1_nil V). rewrite -> (S_list_copy_2_nil V). reflexivity. - rewrite -> (S_list_copy_1_cons V v vs'). rewrite -> (S_list_copy_2_cons V v vs'). rewrite -> IHvs'. reflexivity. Qed. (* {END} *) (* Since our list_copy_identity implementation satisfies the specification_of_list_copy, and we have proven that there_is_at_most_one_list_copy_function, we can conclude that list_copy and list_copy_identity are equivalent. *) (* {task_3_g2} *) Corollary list_copy_is_equivalant_to_list_copy_identity : forall (V : Type) (vs : list V), list_copy V vs = list_copy_identity V vs. Proof. intros V vs. exact (there_is_at_most_one_list_copy_function V list_copy list_copy_identity list_copy_satisfies_the_specification_of_list_copy list_copy_identity_satisfies_the_specification_of_list_copy vs). Qed. (* {END} *) (* ********** *) (* A study of the polymorphic append function: *) Definition specification_of_list_append (append : forall V : Type, list V -> list V -> list V) := (forall (V : Type) (v2s : list V), append V nil v2s = v2s) /\ (forall (V : Type) (v1 : V) (v1s' v2s : list V), append V (v1 :: v1s') v2s = v1 :: append V v1s' v2s). (* ***** *) (* Task 4: *) (* a. Define a unit-test function for list_append. *) (* {task_4_a} *) Definition test_list_append (candidate : forall V : Type, list V -> list V -> list V) := (eqb_list nat Nat.eqb (candidate nat nil nil) nil) && (eqb_list bool Bool.eqb (candidate bool nil nil) nil) && (eqb_list nat Nat.eqb (candidate nat (1 :: nil) nil) (1 :: nil)) && (eqb_list bool Bool.eqb (candidate bool (true :: nil) nil) (true :: nil)) && (eqb_list nat Nat.eqb (candidate nat (3 :: 2 :: nil) (1 :: nil)) (3 :: 2 :: 1 :: nil)) && (eqb_list bool Bool.eqb (candidate bool (false :: true :: nil) (false :: true :: nil)) (false :: true :: false :: true :: nil)) && (eqb_list (list nat) (eqb_list nat Nat.eqb) (candidate (list nat) ((1 :: nil) :: (2 :: nil) :: nil) ((3 :: 4 :: nil) :: nil)) ((1 :: nil) :: (2 :: nil) :: (3 :: 4 :: nil) :: nil)) && (eqb_list (list bool) (eqb_list bool Bool.eqb) (candidate (list bool) ((false :: nil) :: (true :: nil) :: nil) ((false :: true :: nil) :: nil)) ((false :: nil) :: (true :: nil) :: (false :: true :: nil) :: nil)). (* {END} *) (* b. Implement the list_append function recursively. *) (* {task_4_b} *) Fixpoint list_append (V : Type) (v1s v2s : list V) : list V := match v1s with nil => v2s | v1 :: v1s' => v1 :: (list_append V v1s' v2s) end. Compute (test_list_append list_append). (* {END} *) (* c. State its associated fold-unfold lemmas. *) (* {task_4_c} *) Lemma fold_unfold_list_append_nil : forall (V : Type) (v2s : list V), list_append V nil v2s = v2s. Proof. fold_unfold_tactic list_append. Qed. Lemma fold_unfold_list_append_cons : forall (V : Type) (v1 : V) (v1s' v2s : list V), list_append V (v1 :: v1s') v2s = v1 :: (list_append V v1s' v2s). Proof. fold_unfold_tactic list_append. Qed. (* {END} *) (* d. Prove that your implementation satisfies the specification. *) (* {task_4_d} *) Theorem list_append_satisfies_the_specification_of_list_append : specification_of_list_append list_append. Proof. unfold specification_of_list_append. exact (conj fold_unfold_list_append_nil fold_unfold_list_append_cons). Qed. (* {END} *) (* {task_4_d2} *) Theorem there_is_at_most_one_list_append_function : forall (V : Type) (list_append_1 list_append_2 : forall V : Type, list V -> list V -> list V), specification_of_list_append list_append_1 -> specification_of_list_append list_append_2 -> forall v1s v2s : list V, list_append_1 V v1s v2s = list_append_2 V v1s v2s. Proof. intros V list_append_1 list_append_2. unfold specification_of_list_copy. intros [S_list_append_1_nil S_list_append_1_cons] [S_list_append_2_nil S_list_append_2_cons] v1s v2s. induction v1s as [ | v1 v1s' IHv1s']. - rewrite -> (S_list_append_1_nil V v2s). rewrite -> (S_list_append_2_nil V v2s). reflexivity. - rewrite -> (S_list_append_1_cons V v1 v1s' v2s). rewrite -> (S_list_append_2_cons V v1 v1s' v2s). rewrite -> IHv1s'. reflexivity. Qed. (* {END} *) (* e. Prove whether nil is neutral on the left of list_append. *) (* {task_4_e} *) Property nil_is_left_neutral_wrt_list_append : forall (V : Type) (v2s : list V), list_append V nil v2s = v2s. Proof. exact fold_unfold_list_append_nil. Qed. (* {END} *) (* f. Prove whether nil is neutral on the right of list_append. *) (* {task_4_f} *) Property nil_is_right_neutral_wrt_list_append : forall (V : Type) (v1s : list V), list_append V v1s nil = v1s. Proof. intros V v1s. induction v1s as [ | v1 v1s' IHv1s']. - exact (fold_unfold_list_append_nil V nil). - rewrite -> (fold_unfold_list_append_cons V v1 v1s'). rewrite -> IHv1s'. reflexivity. Qed. (* {END} *) (* g. Prove whether list_append is commutative. *) (* {task_4_g} *) Property list_append_is_not_commutative : exists (V : Type) (v1s v2s : list V), list_append V v1s v2s <> list_append V v2s v1s. Proof. exists bool. exists (false :: nil). exists (true :: nil). rewrite -> (fold_unfold_list_append_cons bool false nil). rewrite -> (fold_unfold_list_append_nil bool). rewrite -> (fold_unfold_list_append_cons bool true nil). rewrite -> (fold_unfold_list_append_nil bool). unfold not. intro H_absurd. discriminate H_absurd. Qed. (* {END} *) (* h. Prove whether list_append is associative. *) (* {task_4_h} *) Property list_append_is_associative : forall (V : Type) (v1s v2s v3s : list V), list_append V v1s (list_append V v2s v3s) = list_append V (list_append V v1s v2s) v3s. Proof. intros V v1s v2s v3s. induction v1s as [ | v1 v1s']. - rewrite -> (fold_unfold_list_append_nil V (list_append V v2s v3s)). rewrite -> (fold_unfold_list_append_nil V v2s). reflexivity. - rewrite -> (fold_unfold_list_append_cons V v1 v1s' (list_append V v2s v3s)). rewrite -> (fold_unfold_list_append_cons V v1 v1s' v2s). rewrite -> (fold_unfold_list_append_cons V v1 (list_append V v1s' v2s) v3s). rewrite -> IHv1s'. reflexivity. Qed. (* {END} *) (* i. Prove whether appending two lists preserves their length. *) (* {task_4_i} *) Proposition list_append_and_list_length_commute_with_each_other : forall (V : Type) (v1s v2s : list V), list_length V (list_append V v1s v2s) = list_length V v1s + list_length V v2s. Proof. intros V v1s v2s. induction v1s as [ | v1 v1s' IHv1s']. - rewrite -> (fold_unfold_list_append_nil V v2s). rewrite -> (fold_unfold_list_length_nil V). rewrite -> (Nat.add_0_l (list_length V v2s)). reflexivity. - rewrite -> (fold_unfold_list_append_cons V v1 v1s'). rewrite -> (fold_unfold_list_length_cons V v1 (list_append V v1s' v2s)). rewrite -> (fold_unfold_list_length_cons V v1 v1s'). rewrite -> IHv1s'. rewrite -> (Nat.add_succ_l (list_length V v1s') (list_length V v2s)). reflexivity. Qed. (* {END} *) (* j. Prove whether list_append and list_copy commute with each other. *) (* {task_4_j} *) Proposition list_append_and_list_copy_commute_with_each_other : forall (V : Type) (v1s v2s : list V), list_copy V (list_append V v1s v2s) = list_append V (list_copy V v1s) (list_copy V v2s). Proof. intros V v1s v2s. induction v1s as [ | v1 v1s' IHv1s']. - rewrite -> (fold_unfold_list_append_nil V v2s). rewrite -> (fold_unfold_list_copy_nil V). rewrite -> (fold_unfold_list_append_nil V (list_copy V v2s)). reflexivity. - rewrite -> (fold_unfold_list_append_cons V v1 v1s'). rewrite -> (fold_unfold_list_copy_cons V v1 (list_append V v1s' v2s)). rewrite -> (fold_unfold_list_copy_cons V v1 v1s'). rewrite -> IHv1s'. rewrite -> (fold_unfold_list_append_cons V v1 (list_copy V v1s')). reflexivity. Restart. (* Alternative proof using ideas from Task 3g *) intros V v1s v2s. rewrite -> (list_copy_is_equivalant_to_list_copy_identity V v1s). rewrite -> (list_copy_is_equivalant_to_list_copy_identity V v2s). rewrite -> (list_copy_is_equivalant_to_list_copy_identity V (list_append V v1s v2s)). unfold list_copy_identity. reflexivity. Qed. (* {END} *) (* ********** *) (* A study of the polymorphic reverse function: *) (* {specification_of_list_reverse} *) Definition specification_of_list_reverse (reverse : forall V : Type, list V -> list V) := forall append : forall W : Type, list W -> list W -> list W, specification_of_list_append append -> (forall V : Type, reverse V nil = nil) /\ (forall (V : Type) (v : V) (vs' : list V), reverse V (v :: vs') = append V (reverse V vs') (v :: nil)). (* {END} *) (* ***** *) (* Task 5: *) (* a. Define a unit-test function for an implementation of the reverse function. *) (* {task_5_a} *) Definition test_list_reverse (candidate : forall V : Type, list V -> list V) := (eqb_list nat Nat.eqb (candidate nat nil) nil) && (eqb_list bool Bool.eqb (candidate bool nil) nil) && (eqb_list nat Nat.eqb (candidate nat (1 :: nil)) (1 :: nil)) && (eqb_list bool Bool.eqb (candidate bool (true :: nil)) (true :: nil)) && (eqb_list nat Nat.eqb (candidate nat (2 :: 1 :: nil)) (1 :: 2 :: nil)) && (eqb_list bool Bool.eqb (candidate bool (false :: true :: nil)) (true :: false :: nil)) && (eqb_list nat Nat.eqb (candidate nat (3 :: 2 :: 1 :: nil)) (1 :: 2 :: 3 :: nil)) && (eqb_list bool Bool.eqb (candidate bool (true :: false :: true :: nil)) (true :: false :: true :: nil)) && (eqb_list (list nat) (eqb_list nat Nat.eqb) (candidate (list nat) ((1 :: 2 :: nil) :: (2 :: 3 :: nil) :: nil)) ((2 :: 3 :: nil) :: (1 :: 2 :: nil) :: nil)) && (eqb_list (list bool) (eqb_list bool Bool.eqb) (candidate (list bool) ((true :: false :: nil) :: (true :: nil) :: nil)) ((true :: nil) :: (true :: false :: nil) :: nil)). (* {END} *) (* b. Implement the reverse function recursively, using list_append. *) (* {task_5_b} *) Fixpoint list_reverse (V : Type) (vs : list V) : list V := match vs with nil => nil | v :: vs' => list_append V (list_reverse V vs') (v :: nil) end. Compute (test_list_reverse list_reverse). (* {END} *) (* c. State the associated fold-unfold lemmas. *) (* {task_5_c} *) Lemma fold_unfold_list_reverse_nil : forall (V : Type), list_reverse V nil = nil. Proof. fold_unfold_tactic list_reverse. Qed. Lemma fold_unfold_list_reverse_cons : forall (V : Type) (v : V) (vs' : list V), list_reverse V (v :: vs') = list_append V (list_reverse V vs') (v :: nil). Proof. fold_unfold_tactic list_reverse. Qed. (* {END} *) (* d. Prove whether your implementation satisfies the specification. *) (* {task_5_d} *) Theorem list_reverse_satisfies_the_specification_of_list_reverse : specification_of_list_reverse list_reverse. Proof. unfold specification_of_list_reverse. intros append S_append. split. - exact fold_unfold_list_reverse_nil. - intros V v vs'. rewrite -> (fold_unfold_list_reverse_cons V v vs'). exact (there_is_at_most_one_list_append_function V list_append append list_append_satisfies_the_specification_of_list_append S_append (list_reverse V vs') (v :: nil)). Qed. (* {END} *) (* e. Prove whether list_reverse is involutory. *) (* {task_5_e} *) Lemma about_list_reverse_and_list_append : forall (V : Type) (v : V) (vs : list V), list_reverse V (list_append V vs (v :: nil)) = v :: (list_reverse V vs). Proof. intros V v vs. induction vs as [ | v' vs' IHvs']. - rewrite -> (fold_unfold_list_append_nil V (v :: nil)). rewrite -> (fold_unfold_list_reverse_cons V v nil). rewrite -> (fold_unfold_list_reverse_nil V). rewrite -> (fold_unfold_list_append_nil V (v :: nil)). reflexivity. - rewrite -> (fold_unfold_list_append_cons V v' vs'). rewrite -> (fold_unfold_list_reverse_cons V v' (list_append V vs' (v :: nil))). rewrite -> IHvs'. rewrite -> (fold_unfold_list_append_cons V v (list_reverse V vs') (v' :: nil)). rewrite -> (fold_unfold_list_reverse_cons V v' vs'). reflexivity. Qed. Lemma about_list_reverse_and_append_alt : forall append : forall W : Type, list W -> list W -> list W, specification_of_list_append append -> forall (V : Type) (v : V) (vs : list V), list_reverse V (append V vs (v :: nil)) = append V (v :: nil) (list_reverse V vs). Proof. intros append S_append V v vs. unfold specification_of_list_append in S_append. destruct S_append as [H_append_nil H_append_cons]. induction vs as [ | v' vs' IHvs']. - rewrite -> (H_append_nil V (v :: nil)). rewrite -> (fold_unfold_list_reverse_nil V). rewrite -> (fold_unfold_list_reverse_cons V v nil). rewrite -> (H_append_cons V v nil nil). rewrite -> (fold_unfold_list_reverse_nil). rewrite -> (H_append_nil V nil). exact (fold_unfold_list_append_nil V (v :: nil)). - rewrite -> (H_append_cons V v' vs' (v :: nil)). rewrite -> (fold_unfold_list_reverse_cons V v' (append V vs' (v :: nil))). rewrite -> IHvs'. rewrite -> (H_append_cons V v nil (list_reverse V (v' :: vs'))). rewrite -> (H_append_cons V v nil (list_reverse V vs')). rewrite -> (H_append_nil V (list_reverse V vs')). rewrite -> (H_append_nil V (list_reverse V (v' :: vs'))). rewrite -> (fold_unfold_list_append_cons V v (list_reverse V vs') (v' :: nil)). rewrite <- (fold_unfold_list_reverse_cons V v' vs'). reflexivity. Qed. (* {END} *) (* {task_5_e2} *) Proposition list_reverse_is_involutory : forall (V : Type) (vs : list V), list_reverse V (list_reverse V vs) = vs. Proof. intros V vs. induction vs as [ | v vs' IHvs']. - rewrite -> (fold_unfold_list_reverse_nil V). rewrite -> (fold_unfold_list_reverse_nil V). reflexivity. - rewrite -> (fold_unfold_list_reverse_cons V v vs'). rewrite -> (about_list_reverse_and_list_append V v (list_reverse V vs')). rewrite -> IHvs'. reflexivity. Restart. (* Using alt eureka lemma *) intros V vs. induction vs as [ | v vs' IHvs']. - rewrite -> (fold_unfold_list_reverse_nil V). rewrite -> (fold_unfold_list_reverse_nil V). reflexivity. - rewrite -> (fold_unfold_list_reverse_cons V v vs'). rewrite -> (about_list_reverse_and_append_alt list_append list_append_satisfies_the_specification_of_list_append V v (list_reverse V vs')). rewrite -> (fold_unfold_list_append_cons V v nil (list_reverse V (list_reverse V vs'))). rewrite -> (fold_unfold_list_append_nil V (list_reverse V (list_reverse V vs'))). rewrite -> IHvs'. reflexivity. Qed. (* {END} *) (* f. Prove whether reversing a list preserves its length. *) (* {task_5_f} *) Proposition list_reverse_and_list_length_commute_with_each_other : forall (V : Type) (vs : list V), list_length V (list_reverse V vs) = list_length V vs. Proof. intros V vs. induction vs as [ | v vs' IHvs']. - rewrite -> (fold_unfold_list_reverse_nil V). reflexivity. - rewrite -> (fold_unfold_list_reverse_cons V v vs'). rewrite -> (list_append_and_list_length_commute_with_each_other V (list_reverse V vs') (v :: nil)). rewrite -> IHvs'. rewrite -> (fold_unfold_list_length_cons V v nil). rewrite -> (fold_unfold_list_length_nil V). rewrite -> (Nat.add_succ_r (list_length V vs')). rewrite -> (fold_unfold_list_length_cons V v vs'). rewrite -> (Nat.add_0_r (list_length V vs')). reflexivity. Qed. (* {END} *) (* g. Do list_append and list_reverse commute with each other (hint: yes they do) and if so how? *) (* {task_5_g} *) Proposition list_reverse_and_list_append_commute_with_each_other : forall (V : Type) (v1s v2s : list V), list_append V (list_reverse V v2s) (list_reverse V v1s) = list_reverse V (list_append V v1s v2s). Proof. intros V v1s v2s. induction v1s as [ | v v1s' IHv1s']. - rewrite -> (fold_unfold_list_reverse_nil V). rewrite -> (nil_is_right_neutral_wrt_list_append V (list_reverse V v2s)). rewrite -> (fold_unfold_list_append_nil V v2s). reflexivity. - rewrite -> (fold_unfold_list_reverse_cons V v v1s'). rewrite -> (fold_unfold_list_append_cons V v v1s'). rewrite -> (fold_unfold_list_reverse_cons V v (list_append V v1s' v2s)). rewrite -> (list_append_is_associative V (list_reverse V v2s) (list_reverse V v1s') (v :: nil)). rewrite -> IHv1s'. reflexivity. Qed. (* {END} *) (* h. Implement the reverse function using an accumulator instead of using list_append. *) (* {task_5_h} *) Fixpoint list_reverse_acc (V : Type) (vs a : list V) : list V := match vs with nil => a | v :: vs' => list_reverse_acc V vs' (v :: a) end. Definition list_reverse_alt (V : Type) (vs : list V) : list V := list_reverse_acc V vs nil. Compute (test_list_reverse list_reverse_alt). (* {END} *) (* i. Revisit the propositions above (involution, preservation of length, commutation with append) and prove whether reverse_v1 satisfies them. Two proof strategies are possible: (1) self-contained proofs with Eureka lemmas, and (2) proofs that hinge on the equivalence of list_reverse_alt and list_reverse. This subtask is very instructive, but optional. *) (* {task_5_i} *) Lemma fold_unfold_list_reverse_acc_nil : forall (V : Type) (acc : list V), list_reverse_acc V nil acc = acc. Proof. fold_unfold_tactic list_reverse_acc. Qed. Lemma fold_unfold_list_reverse_acc_cons : forall (V : Type) (v : V) (vs' acc : list V), list_reverse_acc V (v :: vs') acc = list_reverse_acc V vs' (v :: acc). Proof. fold_unfold_tactic list_reverse. Qed. Lemma about_list_reverse_acc_and_append : forall (V : Type) (vs1 vs2 : list V), list_reverse_acc V vs1 vs2 = list_append V (list_reverse_acc V vs1 nil) vs2. Proof. intros V vs1. induction vs1 as [ | v1 v1s' IHv1s']. - intro vs2. rewrite -> (fold_unfold_list_reverse_acc_nil V vs2). rewrite -> (fold_unfold_list_reverse_acc_nil V nil). exact (fold_unfold_list_append_nil V vs2). - intro vs2. rewrite -> (fold_unfold_list_reverse_acc_cons V v1 v1s'). rewrite -> (fold_unfold_list_reverse_acc_cons V v1 v1s'). rewrite (IHv1s' (v1 :: nil)). rewrite <- (list_append_is_associative V (list_reverse_acc V v1s' nil) (v1 :: nil) vs2). rewrite -> (fold_unfold_list_append_cons V v1 nil vs2). rewrite -> (fold_unfold_list_append_nil V vs2). exact (IHv1s' (v1 :: vs2)). Qed. (* {END} *) (* {task_5_i1} *) Proposition list_reverse_is_equivalent_to_list_reverse_acc : forall (V : Type) (vs : list V), list_reverse V vs = list_reverse_acc V vs nil. Proof. intros V vs. induction vs as [ | v vs' IHvs']. - rewrite -> (fold_unfold_list_reverse_nil V). exact (eq_sym (fold_unfold_list_reverse_acc_nil V nil)). - rewrite -> (fold_unfold_list_reverse_cons V v vs'). rewrite -> (fold_unfold_list_reverse_acc_cons V v vs'). rewrite -> IHvs'. exact (eq_sym (about_list_reverse_acc_and_append V vs' (v :: nil))). Qed. Corollary list_reverse_is_equivalent_to_list_reverse_alt : forall (V : Type) (vs : list V), list_reverse V vs = list_reverse_alt V vs. Proof. unfold list_reverse_alt. exact list_reverse_is_equivalent_to_list_reverse_acc. Qed. (* {END} *) (* {task_5_i2} *) Proposition list_reverse_acc_and_list_append_commute_with_each_other : forall (V : Type) (v1s v2s: list V), list_append V (list_reverse_acc V v2s nil) (list_reverse_acc V v1s nil) = list_reverse_acc V (list_append V v1s v2s) nil. Proof. intros V v1s v2s. induction v1s as [ | v1 v1s' IHv1s']. - rewrite -> (fold_unfold_list_reverse_acc_nil V nil). rewrite -> (nil_is_right_neutral_wrt_list_append V (list_reverse_acc V v2s nil)). rewrite -> (fold_unfold_list_append_nil V v2s). reflexivity. - rewrite -> (fold_unfold_list_reverse_acc_cons V v1 v1s' nil). rewrite -> (fold_unfold_list_append_cons V v1 v1s'). rewrite -> (fold_unfold_list_reverse_acc_cons V v1 (list_append V v1s' v2s) nil). rewrite -> (about_list_reverse_acc_and_append V (list_append V v1s' v2s) (v1 :: nil)). rewrite -> (about_list_reverse_acc_and_append V v1s' (v1 :: nil)). rewrite <- IHv1s'. exact (list_append_is_associative V (list_reverse_acc V v2s nil) (list_reverse_acc V v1s' nil) (v1 :: nil)). Restart. (* Proof by equivalence*) intros V v1s v2s. rewrite <- (list_reverse_is_equivalent_to_list_reverse_acc V v2s). rewrite <- (list_reverse_is_equivalent_to_list_reverse_acc V v1s). rewrite <- (list_reverse_is_equivalent_to_list_reverse_acc V (list_append V v1s v2s)). exact (list_reverse_and_list_append_commute_with_each_other V v1s v2s). Qed. Corollary list_reverse_alt_and_list_append_commute_with_each_other : forall (V : Type) (v1s v2s : list V), list_append V (list_reverse_alt V v2s) (list_reverse_alt V v1s) = list_reverse_alt V (list_append V v1s v2s). Proof. unfold list_reverse_alt. exact list_reverse_acc_and_list_append_commute_with_each_other. Qed. Proposition list_reverse_acc_and_list_length_commute_with_each_other : forall (V : Type) (vs : list V), list_length V (list_reverse_acc V vs nil) = list_length V vs. Proof. intros V vs. induction vs as [ | v vs' IHvs']. - rewrite -> (fold_unfold_list_reverse_acc_nil V nil). reflexivity. - rewrite -> (fold_unfold_list_reverse_acc_cons V v vs' nil). rewrite -> (fold_unfold_list_length_cons V v vs'). rewrite -> (about_list_reverse_acc_and_append V vs' (v :: nil)). rewrite <- IHvs'. rewrite -> (list_append_and_list_length_commute_with_each_other V (list_reverse_acc V vs' nil) (v :: nil)). rewrite -> (fold_unfold_list_length_cons V v nil). rewrite -> (fold_unfold_list_length_nil V). exact (Nat.add_1_r (list_length V (list_reverse_acc V vs' nil))). Restart. (* Proof by equivalence*) intros V vs. rewrite <- (list_reverse_is_equivalent_to_list_reverse_acc V vs). exact (list_reverse_and_list_length_commute_with_each_other V vs). Qed. Corollary list_reverse_alt_and_list_length_commute_with_each_other : forall (V : Type) (vs : list V), list_length V (list_reverse_alt V vs) = list_length V vs. Proof. unfold list_reverse_alt. exact list_reverse_acc_and_list_length_commute_with_each_other. Qed. Proposition list_reverse_acc_is_involutory : forall (V : Type) (vs : list V), list_reverse_acc V (list_reverse_acc V vs nil) nil = vs. Proof. intros V vs. induction vs as [ | v vs' IHvs']. - rewrite -> (fold_unfold_list_reverse_acc_nil V nil). exact (fold_unfold_list_reverse_acc_nil V nil). - rewrite -> (fold_unfold_list_reverse_acc_cons V v vs' nil). rewrite -> (about_list_reverse_acc_and_append V vs' (v :: nil)). rewrite <- (list_reverse_acc_and_list_append_commute_with_each_other V (list_reverse_acc V vs' nil) (v :: nil)). rewrite -> IHvs'. rewrite -> (fold_unfold_list_reverse_acc_cons V v nil nil). rewrite -> (fold_unfold_list_reverse_acc_nil V (v :: nil)). rewrite -> (fold_unfold_list_append_cons V v nil vs'). rewrite -> (fold_unfold_list_append_nil V vs'). reflexivity. Restart. (* Proof by equivalence*) intros V vs. rewrite <- (list_reverse_is_equivalent_to_list_reverse_acc V vs). rewrite <- (list_reverse_is_equivalent_to_list_reverse_acc V (list_reverse V vs)). exact (list_reverse_is_involutory V vs). Qed. Corollary list_reverse_alt_is_involutory : forall (V : Type) (vs : list V), list_reverse_alt V (list_reverse_alt V vs) = vs. Proof. unfold list_reverse_alt. exact list_reverse_acc_is_involutory. Qed. (* {END} *) (* ********** *) (* A study of the polymorphic map function: *) (* {specification_of_list_map} *) Definition specification_of_list_map (map : forall V W : Type, (V -> W) -> list V -> list W) := (forall (V W : Type) (f : V -> W), map V W f nil = nil) /\ (forall (V W : Type) (f : V -> W) (v : V) (vs' : list V), map V W f (v :: vs') = f v :: map V W f vs'). (* {END} *) (* ***** *) (* Task 6: a. Prove whether the specification specifies at most one map function. *) (* {task_6_a} *) Proposition there_is_at_most_one_list_map_function : forall list_map1 list_map2 : forall V W : Type, (V -> W) -> list V -> list W, specification_of_list_map list_map1 -> specification_of_list_map list_map2 -> forall (V W : Type) (f : V -> W) (vs : list V), list_map1 V W f vs = list_map2 V W f vs. Proof. intros list_map1 list_map2 S_list_map1 S_list_map2 V W f vs. induction vs as [ | v vs' IHvs']. - unfold specification_of_list_map in S_list_map1. destruct S_list_map1 as [fold_unfold_list_map1_nil _]. destruct S_list_map2 as [fold_unfold_list_map2_nil _]. rewrite -> (fold_unfold_list_map2_nil V W f). exact (fold_unfold_list_map1_nil V W f). - unfold specification_of_list_map in S_list_map1. destruct S_list_map1 as [_ fold_unfold_list_map1_cons]. destruct S_list_map2 as [_ fold_unfold_list_map2_cons]. rewrite -> (fold_unfold_list_map1_cons V W f v vs'). rewrite -> (fold_unfold_list_map2_cons V W f v vs'). rewrite -> IHvs'. reflexivity. Qed. (* {END} *) (* b. Implement the map function recursively. *) (* {task_6_b} *) Fixpoint list_map (V W : Type) (f : V -> W) (vs : list V) : list W := match vs with nil => nil | v :: vs' => f v :: list_map V W f vs' end. (* c. State the associated fold-unfold lemmas. *) (* {task_6_c} *) Lemma fold_unfold_list_map_nil : forall (V W : Type) (f : V -> W), list_map V W f nil = nil. Proof. fold_unfold_tactic list_map. Qed. Lemma fold_unfold_list_map_cons : forall (V W : Type) (f : V -> W) (v : V) (vs' : list V), list_map V W f (v :: vs') = f v :: list_map V W f vs'. Proof. fold_unfold_tactic list_map. Qed. (* {END} *) (* d. Prove whether your implementation satisfies the specification. *) (* {task_6_d} *) Proposition list_map_satisfies_the_specification_of_list_map : specification_of_list_map list_map. Proof. unfold specification_of_list_map. exact (conj fold_unfold_list_map_nil fold_unfold_list_map_cons). Qed. (* {END} *) (* e. Implement the copy function as an instance of list_map. *) (* {task_6_e} *) Definition identity (V : Type) (v : V) : V := v. Definition list_copy_as_list_map (V : Type) (vs : list V) : list V := list_map V V (identity V) vs. (* Hint: Does list_copy_as_list_map satisfy the specification of list_copy? *) Theorem list_copy_as_list_map_satisfies_the_specification_of_list_copy : specification_of_list_copy list_copy_as_list_map. Proof. unfold specification_of_list_copy. split. - intro V. unfold list_copy_as_list_map. rewrite -> (fold_unfold_list_map_nil V V (identity V)). reflexivity. - intros V v vs'. unfold list_copy_as_list_map. rewrite -> (fold_unfold_list_map_cons V V (identity V) v vs'). unfold identity at 1. reflexivity. Qed. (* {END} *) (* f. Prove whether mapping a function over a list preserves the length of this list. *) (* {task_6_f} *) Proposition list_map_and_list_length_commute_with_each_other : forall (V W : Type) (f : V -> W) (vs : list V), list_length W (list_map V W f vs) = list_length V vs. Proof. intros V W f vs. induction vs as [ | v vs' IHvs']. - rewrite -> (fold_unfold_list_map_nil V W f). rewrite -> (fold_unfold_list_length_nil W). rewrite -> (fold_unfold_list_length_nil V). reflexivity. - rewrite -> (fold_unfold_list_map_cons V W f v vs'). rewrite -> (fold_unfold_list_length_cons W (f v) (list_map V W f vs')). rewrite -> (fold_unfold_list_length_cons V v vs'). rewrite -> IHvs'. reflexivity. Qed. (* {END} *) (* g. Do list_map and list_append commute with each other and if so how? *) (* {task_6_g} *) Proposition list_map_and_list_append_commute_with_each_other : forall (V W : Type) (f : V -> W) (v1s v2s : list V), list_append W (list_map V W f v1s) (list_map V W f v2s) = list_map V W f (list_append V v1s v2s). Proof. intros V W f v1s v2s. induction v1s as [ | v v1s' IHv1s']. - rewrite -> (fold_unfold_list_append_nil V v2s). rewrite -> (fold_unfold_list_map_nil V W f). rewrite -> (fold_unfold_list_append_nil W (list_map V W f v2s)). reflexivity. - rewrite -> (fold_unfold_list_map_cons V W f v v1s'). rewrite -> (fold_unfold_list_append_cons W (f v) (list_map V W f v1s') (list_map V W f v2s)). rewrite -> IHv1s'. rewrite -> (fold_unfold_list_append_cons V v v1s' v2s). rewrite -> (fold_unfold_list_map_cons V W f v (list_append V v1s' v2s)). reflexivity. Qed. (* {END} *) (* h. Do list_map and list_reverse commute with each other and if so how? *) (* {task_6_h} *) Proposition list_map_and_list_reverse_commute_with_each_other : forall (V W : Type) (f : V -> W) (vs : list V), list_reverse W (list_map V W f vs) = list_map V W f (list_reverse V vs). Proof. intros V W f vs. induction vs as [ | v vs' IHvs']. - rewrite -> (fold_unfold_list_map_nil V W f). rewrite -> (fold_unfold_list_reverse_nil W). rewrite -> (fold_unfold_list_reverse_nil V). rewrite -> (fold_unfold_list_map_nil V W f). reflexivity. - rewrite -> (fold_unfold_list_map_cons V W f v vs'). rewrite -> (fold_unfold_list_reverse_cons W (f v) (list_map V W f vs')). rewrite -> IHvs'. rewrite -> (fold_unfold_list_reverse_cons V v vs'). rewrite <- (fold_unfold_list_map_nil V W f). rewrite <- (fold_unfold_list_map_cons V W f v nil). rewrite -> (list_map_and_list_append_commute_with_each_other V W f (list_reverse V vs') (v :: nil)). reflexivity. Qed. (* {END} *) (* i. Do list_map and list_reverse_alt commute with each other and if so how? *) (* {task_6_i} *) Proposition list_map_and_list_reverse_alt_commute_with_each_other : forall (V W : Type) (f : V -> W) (vs : list V), list_reverse_alt W (list_map V W f vs) = list_map V W f (list_reverse_alt V vs). Proof. intros V W f vs. Check list_reverse_is_equivalent_to_list_reverse_alt. rewrite <- (list_reverse_is_equivalent_to_list_reverse_alt W (list_map V W f vs)). rewrite <- (list_reverse_is_equivalent_to_list_reverse_alt V vs). exact (list_map_and_list_reverse_commute_with_each_other V W f vs). Qed. (* {END} *) (* j. Define a unit-test function for the map function and verify that your implementation satisfies it. *) (* {task_6_j} *) Fixpoint evenp (n : nat): bool := match n with | 0 => true | S n' => match n' with | 0 => false | S n'' => evenp n'' end end. Definition test_list_map (candidate : forall V W : Type, (V -> W) -> list V -> list W) := (eqb_list nat Nat.eqb (candidate nat nat (fun v => v) nil) nil) && (eqb_list nat Nat.eqb (candidate nat nat (fun v => v) (1 :: 2:: 3 :: nil)) (1 :: 2 :: 3 :: nil)) && (eqb_list nat Nat.eqb (candidate nat nat (fun v => S v) (1 :: 2:: 3 :: nil)) (2 :: 3 :: 4 :: nil)) && (eqb_list bool Bool.eqb (candidate nat bool evenp (1 :: 2:: 3 :: nil)) (false :: true :: false :: nil)). Compute test_list_map list_map. (* {END} *) (* ********** *) (* A study of the polymorphic fold-right and fold-left functions: *) (* {specification_of_list_fold_right} *) Definition specification_of_list_fold_right (fold_right : forall V W : Type, W -> (V -> W -> W) -> list V -> W) := (forall (V W : Type) (nil_case : W) (cons_case : V -> W -> W), fold_right V W nil_case cons_case nil = nil_case) /\ (forall (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (v : V) (vs' : list V), fold_right V W nil_case cons_case (v :: vs') = cons_case v (fold_right V W nil_case cons_case vs')). Definition specification_of_list_fold_left (fold_left : forall V W : Type, W -> (V -> W -> W) -> list V -> W) := (forall (V W : Type) (nil_case : W) (cons_case : V -> W -> W), fold_left V W nil_case cons_case nil = nil_case) /\ (forall (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (v : V) (vs' : list V), fold_left V W nil_case cons_case (v :: vs') = fold_left V W (cons_case v nil_case) cons_case vs'). (* {END} *) (* ***** *) (* Task 7: a. Implement the fold-right function recursively. *) (* {task_7_a} *) Fixpoint list_fold_right (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (vs : list V) : W := match vs with nil => nil_case | v :: vs' => cons_case v (list_fold_right V W nil_case cons_case vs') end. (* {END} *) (* b. Implement the fold-left function recursively. *) (* {task_7_b} *) Fixpoint list_fold_left (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (vs : list V) : W := match vs with nil => nil_case | v :: vs' => list_fold_left V W (cons_case v nil_case) cons_case vs' end. (* {END} *) (* c. state the fold-unfold lemmas associated to list_fold_right and to list_fold_left *) (* {task_7_c} *) Lemma fold_unfold_list_fold_right_nil : forall (V W : Type) (nil_case : W) (cons_case : V -> W -> W), list_fold_right V W nil_case cons_case nil = nil_case. Proof. fold_unfold_tactic list_fold_right. Qed. Lemma fold_unfold_list_fold_right_cons : forall (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (v : V) (vs' : list V), list_fold_right V W nil_case cons_case (v :: vs') = cons_case v (list_fold_right V W nil_case cons_case vs'). Proof. fold_unfold_tactic list_fold_right. Qed. Lemma fold_unfold_list_fold_left_nil : forall (V W : Type) (nil_case : W) (cons_case : V -> W -> W), list_fold_left V W nil_case cons_case nil = nil_case. Proof. fold_unfold_tactic list_fold_left. Qed. Lemma fold_unfold_list_fold_left_cons : forall (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (v : V) (vs' : list V), list_fold_left V W nil_case cons_case (v :: vs') = list_fold_left V W (cons_case v nil_case) cons_case vs'. Proof. fold_unfold_tactic list_fold_left. Qed. (* {END} *) (* d. Prove that each of your implementations satisfies the corresponding specification. *) (* {task_7_d} *) Theorem list_fold_left_satisfies_the_specification_of_fold_left : specification_of_list_fold_left list_fold_left. Proof. unfold specification_of_list_fold_left. exact (conj fold_unfold_list_fold_left_nil fold_unfold_list_fold_left_cons). Qed. Theorem list_fold_right_satisfies_the_specification_of_fold_right : specification_of_list_fold_right list_fold_right. Proof. unfold specification_of_list_fold_right. exact (conj fold_unfold_list_fold_right_nil fold_unfold_list_fold_right_cons). Qed. (* {END} *) (* e. Which function do foo and bar (defined just below) compute? *) (* {task_7_e1} *) Definition foo (V : Type) (vs : list V) := list_fold_right V (list V) nil (fun v vs => v :: vs) vs. Compute test_list_copy foo. Proposition foo_satisfies_the_specification_of_list_copy : specification_of_list_copy foo. Proof. unfold foo. unfold specification_of_list_copy. split. - intro V. remember (fun (v : V) (vs0 : list V) => v :: vs0) as f eqn:H_f. exact (fold_unfold_list_fold_right_nil V (list V) nil f). - intros V v vs'. remember (fun (v : V) (vs0 : list V) => v :: vs0) as f eqn:H_f. rewrite -> H_f. rewrite <- H_f. rewrite -> (fold_unfold_list_fold_right_cons V (list V) nil f v vs'). rewrite -> H_f. reflexivity. Qed. (* {END} *) (* {task_7_e2} *) Definition bar (V : Type) (vs : list V) := list_fold_left V (list V) nil (fun v vs => v :: vs) vs. Compute test_list_reverse bar. Lemma about_bar_and_append : forall (V : Type) (v1s v2s : list V) (f : V -> list V -> list V) (append : forall W : Type, list W -> list W -> list W), specification_of_list_append append -> f = (fun (v : V) (vs : list V) => v :: vs) -> list_fold_left V (list V) v2s f v1s = append V (list_fold_left V (list V) nil f v1s) v2s. Proof. intros V v1s v2s f append S_append H_f. rewrite -> (there_is_at_most_one_list_append_function V append list_append S_append list_append_satisfies_the_specification_of_list_append (list_fold_left V (list V) nil f v1s) v2s). revert v2s. induction v1s as [ | v1 v1s' IHv1s']. - intro v2s. rewrite -> (fold_unfold_list_fold_left_nil V (list V) nil f). rewrite -> (fold_unfold_list_append_nil V v2s). exact (fold_unfold_list_fold_left_nil V (list V) v2s f). - intro v2s. rewrite -> (fold_unfold_list_fold_left_cons V (list V) v2s f v1 v1s'). rewrite -> (fold_unfold_list_fold_left_cons V (list V) nil f v1 v1s'). rewrite -> H_f. rewrite <- H_f. rewrite -> (IHv1s' (v1 :: v2s)). rewrite -> (IHv1s' (v1 :: nil)). rewrite <- (list_append_is_associative V (list_fold_left V (list V) nil f v1s') (v1 :: nil) v2s). rewrite -> (fold_unfold_list_append_cons V v1 nil). rewrite -> (fold_unfold_list_append_nil V v2s). reflexivity. Qed. Proposition bar_satisfies_the_specification_of_list_reverse : specification_of_list_reverse bar. Proof. unfold bar. unfold specification_of_list_reverse. intros append S_append. split. - intro V. remember (fun (v : V) (vs : list V) => v :: vs) as f eqn:H_f. exact (fold_unfold_list_fold_left_nil V (list V) nil f). - intros V v vs'. remember (fun (v : V) (vs : list V) => v :: vs) as f eqn:H_f. rewrite -> H_f. rewrite <- H_f. rewrite -> (fold_unfold_list_fold_left_cons V (list V) nil f v vs'). rewrite -> H_f. rewrite <- H_f. exact (about_bar_and_append V vs' (v :: nil) f append S_append H_f). Qed. (* {END} *) (* f. Implement the length function either as an instance of list_fold_right or as an instance of list_fold_left, and justify your choice. *) (* {task_7_f} *) Definition list_length_using_list_fold_right (V : Type) (vs : list V) : nat := list_fold_right V nat 0 (fun _ len' => S len') vs. Compute test_list_length list_length_using_list_fold_right. Proposition list_length_using_list_fold_left_satisfies_specification : specification_of_list_length list_length_using_list_fold_right. Proof. unfold specification_of_list_length. unfold list_length_using_list_fold_right. split. - intro V. remember (fun (_ : V) (len' : nat) => S len') as f eqn:H_f. exact (fold_unfold_list_fold_right_nil V nat 0 f). - intros V v vs'. remember (fun (_ : V) (len' : nat) => S len') as f eqn:H_f. rewrite -> (fold_unfold_list_fold_right_cons V nat 0 f v vs'). rewrite -> H_f. reflexivity. Qed. (* {END} *) (* g. Implement the copy function either as an instance of list_fold_right or as an instance of list_fold_left, and justify your choice. *) (* {task_7_g} *) Definition list_copy_using_list_fold_right (V : Type) (vs : list V) : list V := list_fold_right V (list V) nil (fun v vs' => v :: vs') vs. Compute test_list_copy list_copy_using_list_fold_right. Proposition list_copy_using_list_fold_right_is_equivalent_to_foo : list_copy_using_list_fold_right = foo. Proof. unfold list_copy_using_list_fold_right. unfold foo. reflexivity. Qed. Corollary list_copy_using_list_fold_right_satisfies_specification : specification_of_list_copy list_copy_using_list_fold_right. Proof. rewrite -> list_copy_using_list_fold_right_is_equivalent_to_foo. exact foo_satisfies_the_specification_of_list_copy. Qed. (* {END} *) (* h. Implement the append function either as an instance of list_fold_right or as an instance of list_fold_left, and justify your choice. *) (* {task_7_h} *) Definition list_append_using_list_fold_right (V : Type) (v1s v2s : list V) : list V := list_fold_right V (list V) v2s (fun v vs => v :: vs) v1s. Compute test_list_append list_append_using_list_fold_right. Proposition list_append_using_list_fold_right_satisfies_specification : specification_of_list_append list_append_using_list_fold_right. Proof. unfold specification_of_list_append. unfold list_append_using_list_fold_right. split. - intros V v2s. remember (fun (v : V) (vs : list V) => v :: vs) as f eqn:H_f. exact (fold_unfold_list_fold_right_nil V (list V) v2s f). - intros V v1 v1s' v2s. remember (fun (v : V) (vs : list V) => v :: vs) as f eqn:H_f. rewrite -> H_f. rewrite <- H_f. rewrite -> (fold_unfold_list_fold_right_cons V (list V) v2s f v1 v1s'). rewrite -> H_f. reflexivity. Qed. (* {END} *) (* i. Implement the reverse function either as an instance of list_fold_right or as an instance of list_fold_left, and justify your choice. *) (* {task_7_i} *) Definition list_reverse_using_list_fold_left (V : Type) (vs : list V) : list V := list_fold_left V (list V) nil (fun v vs => v :: vs) vs. Compute test_list_reverse list_reverse_using_list_fold_left. Proposition list_reverse_using_list_fold_left_is_equivalent_to_bar : list_reverse_using_list_fold_left = bar. Proof. unfold list_reverse_using_list_fold_left. unfold bar. reflexivity. Qed. Corollary list_reverse_using_list_fold_left_satisfies_specification : specification_of_list_reverse list_reverse_using_list_fold_left. Proof. rewrite -> list_reverse_using_list_fold_left_is_equivalent_to_bar. exact bar_satisfies_the_specification_of_list_reverse. Qed. (* {END} *) (* j. Implement the map function either as an instance of list_fold_right or as an instance of list_fold_left, and justify your choice. *) (* {task_7_j} *) Definition list_map_using_list_fold_right (V W : Type) (f : V -> W) (vs : list V) : list W := list_fold_right V (list W) nil (fun v ws => f v :: ws) vs. Compute test_list_map list_map_using_list_fold_right. Proposition list_map_using_list_fold_right_satisfies_specification : specification_of_list_map list_map_using_list_fold_right. Proof. unfold specification_of_list_map. unfold list_map_using_list_fold_right. split. - intros V W f. remember (fun (v : V) (ws : list W) => f v :: ws) as cons eqn:H_cons. exact (fold_unfold_list_fold_right_nil V (list W) nil cons). - intros V W f v vs'. remember (fun (v : V) (ws : list W) => f v :: ws) as cons eqn:H_cons. rewrite -> (fold_unfold_list_fold_right_cons V (list W) nil cons v vs'). rewrite -> H_cons. reflexivity. Qed. (* {END} *) (* k. Implement eqb_list either as an instance of list_fold_right or as an instance of list_fold_left, and justify your choice. *) (* {task_7_k} *) Definition eqb_list_using_list_fold_left_cons_case (V : Type) (eqb_V: V -> V -> bool) (v1 : V) (acc : bool * list V) : (bool * list V) := let (ih, v2s) := acc in match v2s with nil => (false, nil) | v2 :: v2s' => (eqb_V v1 v2 && ih, v2s') end. Definition eqb_list_using_list_fold_left (V : Type) (eqb_V: V -> V -> bool) (v1s v2s : list V) : bool := let (ih, v2s') := (list_fold_left V (bool * list V) (true, v2s) (eqb_list_using_list_fold_left_cons_case V eqb_V) v1s) in match v2s' with nil => true | _ :: _ => false end && ih. Check (fold_unfold_eqb_list_cons). (* {END} *) (* {task_7_k_lemma} *) Lemma about_eqb_list_using_list_fold_left : forall (V : Type) (eqb_V : V -> V -> bool) (b : bool) (v1s v2s : list V), (let (ih, v2s'0) := (list_fold_left V (bool * list V) (b, v2s) (eqb_list_using_list_fold_left_cons_case V eqb_V) v1s) in match v2s'0 with | nil => true | _ :: _ => false end && ih) = b && (let (ih, v2s'0) := (list_fold_left V (bool * list V) (true, v2s) (eqb_list_using_list_fold_left_cons_case V eqb_V) v1s) in match v2s'0 with | nil => true | _ :: _ => false end && ih). Proof. intros V eqb_V b v1s. revert b. induction v1s as [ | v1 v1s' IHv1s']. - intros b v2s. rewrite -> (fold_unfold_list_fold_left_nil V (bool * list V) (b, v2s) (eqb_list_using_list_fold_left_cons_case V eqb_V)). rewrite -> (fold_unfold_list_fold_left_nil V (bool * list V) (true, v2s) (eqb_list_using_list_fold_left_cons_case V eqb_V)). rewrite -> (andb_true_r (match v2s with | nil => true | _ :: _ => false end)). exact (andb_comm (match v2s with | nil => true | _ :: _ => false end) b). - intros b v2s. case v2s as [ | v2 v2s'] eqn:H_v2s. + rewrite -> (fold_unfold_list_fold_left_cons V (bool * list V) (b, nil) (eqb_list_using_list_fold_left_cons_case V eqb_V) v1 v1s'). unfold eqb_list_using_list_fold_left_cons_case at 1. rewrite -> (IHv1s' false nil). rewrite -> (fold_unfold_list_fold_left_cons V (bool * list V) (true, nil) (eqb_list_using_list_fold_left_cons_case V eqb_V) v1 v1s'). unfold eqb_list_using_list_fold_left_cons_case at 2. rewrite -> (IHv1s' false nil). rewrite -> (andb_false_l (let (ih, v2s'0) := list_fold_left V (bool * list V) (true, nil) (eqb_list_using_list_fold_left_cons_case V eqb_V) v1s' in match v2s'0 with | nil => true | _ :: _ => false end && ih)). exact (eq_sym (andb_false_r b)). + rewrite -> (fold_unfold_list_fold_left_cons V (bool * list V) (b, v2 :: v2s') (eqb_list_using_list_fold_left_cons_case V eqb_V) v1 v1s'). unfold eqb_list_using_list_fold_left_cons_case at 1. rewrite -> (IHv1s' (eqb_V v1 v2 && b) v2s'). rewrite -> (fold_unfold_list_fold_left_cons V (bool * list V) (true, v2 :: v2s') (eqb_list_using_list_fold_left_cons_case V eqb_V) v1 v1s'). unfold eqb_list_using_list_fold_left_cons_case at 2. rewrite -> (andb_true_r (eqb_V v1 v2)). rewrite -> (IHv1s' (eqb_V v1 v2) v2s'). rewrite -> (andb_comm (eqb_V v1 v2) b). exact (eq_sym (andb_assoc b (eqb_V v1 v2) (let (ih , v2s'0) := list_fold_left V (bool * list V) (true, v2s') (eqb_list_using_list_fold_left_cons_case V eqb_V) v1s' in match v2s'0 with | nil => true | _ :: _ => false end && ih))). Qed. (* {END} *) (* {task_7_k_equiv} *) Proposition eqb_list_using_list_fold_left_is_equivalent_to_eqb_list : forall (V : Type) (v1s v2s : list V) (eqb_V : V -> V -> bool), eqb_list V eqb_V v1s v2s = eqb_list_using_list_fold_left V eqb_V v1s v2s. Proof. intros V v1s v2s eqb_V. revert v2s. induction v1s as [ | v1 v1s' IHv1s']. - intro v2s. unfold eqb_list_using_list_fold_left. rewrite -> (fold_unfold_list_fold_left_nil V (bool * list V) (true, v2s) (eqb_list_using_list_fold_left_cons_case V eqb_V)). rewrite -> (andb_true_r (match v2s with | nil => true | _ :: _ => false end)). exact (fold_unfold_eqb_list_nil V eqb_V v2s). - intro v2s. case v2s as [ | v2 v2s'] eqn:H_v2s. + rewrite -> (fold_unfold_eqb_list_cons V eqb_V v1 v1s' nil). unfold eqb_list_using_list_fold_left. rewrite -> (fold_unfold_list_fold_left_cons V (bool * list V) (true, nil) (eqb_list_using_list_fold_left_cons_case V eqb_V) v1 v1s'). unfold eqb_list_using_list_fold_left_cons_case at 1. rewrite -> (about_eqb_list_using_list_fold_left V eqb_V false v1s' nil). exact (eq_sym (andb_false_l (let (ih, v2s'0) := list_fold_left V (bool * list V) (true, nil) (eqb_list_using_list_fold_left_cons_case V eqb_V) v1s' in match v2s'0 with | nil => true | _ :: _ => false end && ih))). + rewrite -> (fold_unfold_eqb_list_cons V eqb_V v1 v1s' (v2 :: v2s')). unfold eqb_list_using_list_fold_left. rewrite -> (fold_unfold_list_fold_left_cons V (bool * list V) (true, (v2 :: v2s')) (eqb_list_using_list_fold_left_cons_case V eqb_V) v1 v1s'). unfold eqb_list_using_list_fold_left_cons_case at 1. rewrite -> (IHv1s' v2s'). unfold eqb_list_using_list_fold_left. rewrite -> (andb_true_r (eqb_V v1 v2)). exact (eq_sym (about_eqb_list_using_list_fold_left V eqb_V (eqb_V v1 v2) v1s' v2s')). Qed. (* {END} *) (* l. Implement list_fold_right as an instance of list_fold_left, using list_reverse. *) (* {task_7_l} *) Definition list_fold_right_left_using_list_reverse (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (vs : list V) : W := list_fold_left V W nil_case cons_case (list_reverse V vs). (* {END} *) (* {task_7_l_lemma} *) Lemma about_list_fold_left_and_list_append : forall (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (v1s v2s : list V), list_fold_left V W nil_case cons_case (list_append V v1s v2s) = list_fold_left V W (list_fold_left V W nil_case cons_case v1s) cons_case v2s. Proof. intros V W nil_case cons_case v1s. revert nil_case. induction v1s as [ | v1 v1s' IHv1s']. - intros nil_case v2s. rewrite -> (fold_unfold_list_append_nil V v2s). rewrite -> (fold_unfold_list_fold_left_nil V W nil_case cons_case). reflexivity. - intros nil_case v2s. rewrite -> (fold_unfold_list_append_cons V v1 v1s'). rewrite -> (fold_unfold_list_fold_left_cons V W nil_case cons_case v1 (list_append V v1s' v2s)). rewrite -> (IHv1s' (cons_case v1 nil_case) v2s). rewrite <- (fold_unfold_list_fold_left_cons V W nil_case cons_case v1 v1s'). reflexivity. Qed. (* {END} *) (* {task_7_l_satisfies} *) Proposition list_fold_right_left_using_list_reverse_satisfies_specification_of_list_fold_right : specification_of_list_fold_right list_fold_right_left_using_list_reverse. Proof. unfold specification_of_list_fold_right, list_fold_right_left_using_list_reverse. split. - intros V W nil_case cons_case. rewrite -> (fold_unfold_list_reverse_nil V). exact (fold_unfold_list_fold_left_nil V W nil_case cons_case). - intros V W nil_case cons_case v vs'. rewrite -> (fold_unfold_list_reverse_cons V v vs'). rewrite -> (about_list_fold_left_and_list_append V W nil_case cons_case (list_reverse V vs') (v :: nil)). rewrite -> (fold_unfold_list_fold_left_cons V W (list_fold_left V W nil_case cons_case (list_reverse V vs')) cons_case v nil). exact (fold_unfold_list_fold_left_nil V W (cons_case v (list_fold_left V W nil_case cons_case (list_reverse V vs'))) cons_case). Qed. (* {END} *) (* m. Implement list_fold_left as an instance of list_fold_right, using list_reverse. *) (* {task_7_m} *) Definition list_fold_left_right_using_list_reverse (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (vs : list V) : W := list_fold_right V W nil_case cons_case (list_reverse V vs). (* {END} *) (* {task_7_m_lemma} *) Lemma about_list_fold_right_and_list_append : forall (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (v1s v2s : list V), list_fold_right V W nil_case cons_case (list_append V v1s v2s) = list_fold_right V W (list_fold_right V W nil_case cons_case v2s) cons_case v1s. Proof. intros V W nil_case cons_case v1s v2s. induction v1s as [ | v1 v1s' IHv1s']. - rewrite -> (fold_unfold_list_append_nil V v2s). exact (fold_unfold_list_fold_right_nil V W (list_fold_right V W nil_case cons_case v2s) cons_case). - rewrite -> (fold_unfold_list_append_cons V v1 v1s' v2s). rewrite -> (fold_unfold_list_fold_right_cons V W nil_case cons_case v1 (list_append V v1s' v2s)). rewrite -> IHv1s'. rewrite -> (fold_unfold_list_fold_right_cons V W (list_fold_right V W nil_case cons_case v2s) cons_case v1 v1s'). reflexivity. Qed. (* {END} *) (* {task_7_m_satisfies} *) Proposition list_fold_left_right_using_list_reverse_satisfies_specification_of_list_fold_left : specification_of_list_fold_left list_fold_left_right_using_list_reverse. Proof. unfold specification_of_list_fold_left, list_fold_left_right_using_list_reverse. split. - intros V W nil_case cons_case. rewrite -> (fold_unfold_list_reverse_nil V). exact (fold_unfold_list_fold_right_nil V W nil_case cons_case). - intros V W nil_case cons_case v vs'. rewrite -> (fold_unfold_list_reverse_cons V v vs'). rewrite -> (about_list_fold_right_and_list_append V W nil_case cons_case (list_reverse V vs') (v :: nil)). rewrite -> (fold_unfold_list_fold_right_cons V W nil_case cons_case v nil). rewrite -> (fold_unfold_list_fold_right_nil V W nil_case cons_case). reflexivity. Qed. (* {END} *) (* n. Implement list_fold_right as an instance of list_fold_left, without using list_reverse. *) (* {task_7_n} *) Definition list_fold_right_left (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (vs : list V) : W := (list_fold_left V (W -> W) (fun w => w) (fun v stack => fun w => stack (cons_case v w)) vs) nil_case. (* {END} *) (* {task_7_n_lemma} *) Lemma about_list_fold_left_and_function_accumulators : forall (V W : Type) (vs : list V) (nil_case : W) (cons_case : V -> W -> W) (f : W -> W), list_fold_left V (W -> W) (fun w => f w) (fun v stack w => stack (cons_case v w)) vs nil_case = f (list_fold_left V (W -> W) (fun w => w) (fun v stack w => stack (cons_case v w)) vs nil_case). Proof. intros V W vs nil_case cons_case. induction vs as [ | v vs' IHvs']. - intro f. remember (fun (v0 : V) (stack : W -> W) (w : W) => stack (cons_case v0 w)) as push eqn:H_push. rewrite -> (fold_unfold_list_fold_left_nil V (W -> W) (fun w : W => f w) push). rewrite -> (fold_unfold_list_fold_left_nil V (W -> W) (fun w : W => w) push). reflexivity. - intro f. remember (fun (v0 : V) (stack : W -> W) (w : W) => stack (cons_case v0 w)) as push eqn:H_push. rewrite -> (fold_unfold_list_fold_left_cons V (W -> W) (fun w : W => f w) push v vs'). rewrite -> (fold_unfold_list_fold_left_cons V (W -> W) (fun w : W => w) push v vs'). rewrite -> H_push. rewrite <- H_push. rewrite -> (IHvs' (fun w => f (cons_case v w))). rewrite -> (IHvs' (cons_case v)). reflexivity. Qed. (* {END} *) (* {task_7_n_satisfies} *) Proposition list_fold_right_left_satisfies_specification_of_list_fold_right : specification_of_list_fold_right list_fold_right_left. Proof. unfold specification_of_list_fold_right, list_fold_right_left. split. - intros V W nil_case cons_case. remember (fun (v : V) (stack : W -> W) (w : W) => stack (cons_case v w)) as push eqn:H_push. rewrite -> (fold_unfold_list_fold_left_nil V (W -> W) (fun w : W => w) push). reflexivity. - intros V W nil_case cons_case v vs'. remember (fun (v : V) (stack : W -> W) (w : W) => stack (cons_case v w)) as push eqn:H_push. rewrite -> (fold_unfold_list_fold_left_cons V (W -> W) (fun w : W => w) push v vs'). rewrite -> H_push. exact (about_list_fold_left_and_function_accumulators V W vs' nil_case cons_case (fun w : W => cons_case v w)). Qed. (* {END} *) (* o. Implement list_fold_left as an instance of list_fold_right, without using list_reverse. *) (* {task_7_o} *) Definition list_fold_left_right (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (vs : list V) : W := (list_fold_right V (W -> W) (fun w => w) (fun v stack => fun w => stack (cons_case v w)) vs) nil_case. (* {END} *) (* {task_7_o_satisfies} *) Proposition list_fold_left_right_satisfies_specification_of_list_fold_left : specification_of_list_fold_left list_fold_left_right. Proof. unfold specification_of_list_fold_left, list_fold_left_right. split. - intros V W nil_case cons_case. remember (fun (v : V) (stack : W -> W) (w : W) => stack (cons_case v w)) as push eqn:H_push. rewrite -> (fold_unfold_list_fold_right_nil V (W -> W) (fun w : W => w) push). reflexivity. - intros V W nil_case cons_case v vs'. remember (fun (v : V) (stack : W -> W) (w : W) => stack (cons_case v w)) as push eqn:H_push. rewrite -> (fold_unfold_list_fold_right_cons V (W -> W) (fun w : W => w) push v vs'). rewrite -> H_push. reflexivity. Qed. (* {END} *) (* p. Show that if the cons case is a function that is left permutative (defined just below), applying list_fold_left and applying list_fold_right to a nil case, this cons case, and a list give the same result *) (* {task_7_p_definition} *) Definition is_left_permutative (V W : Type) (op2 : V -> W -> W) := forall (v1 v2 : V) (w : W), op2 v1 (op2 v2 w) = op2 v2 (op2 v1 w). (* {END} *) (* {task_7_p_eureka} *) Lemma about_list_fold_left_with_left_permutative_cons_case : forall (V W : Type) (cons_case : V -> W -> W), is_left_permutative V W cons_case -> forall (nil_case : W) (v : V) (vs : list V), list_fold_left V W (cons_case v nil_case) cons_case vs = cons_case v (list_fold_left V W nil_case cons_case vs). Proof. intros V W cons_case. unfold is_left_permutative. intros H_left_permutative nil_case v vs. revert nil_case. induction vs as [ | v' vs' IHvs']. - intro nil_case. rewrite -> (fold_unfold_list_fold_left_nil V W (cons_case v nil_case) cons_case). rewrite -> (fold_unfold_list_fold_left_nil V W nil_case cons_case). reflexivity. - intro nil_case. rewrite -> (fold_unfold_list_fold_left_cons V W (cons_case v nil_case) cons_case v' vs'). rewrite -> (H_left_permutative v' v nil_case). rewrite -> (IHvs' (cons_case v' nil_case)). rewrite -> (fold_unfold_list_fold_left_cons V W nil_case cons_case v' vs'). reflexivity. Qed. (* {END} *) (* {task_7_p} *) Theorem folding_left_and_right_over_lists : forall (V W : Type) (cons_case : V -> W -> W), is_left_permutative V W cons_case -> forall (nil_case : W) (vs : list V), list_fold_left V W nil_case cons_case vs = list_fold_right V W nil_case cons_case vs. Proof. intros V W cons_case. unfold is_left_permutative. intros H_left_permutative nil_case vs. induction vs as [ | v' vs' IHvs']. - rewrite -> (fold_unfold_list_fold_left_nil V W nil_case cons_case). rewrite -> (fold_unfold_list_fold_right_nil V W nil_case cons_case). reflexivity. - rewrite -> (fold_unfold_list_fold_left_cons V W nil_case cons_case v' vs'). rewrite -> (fold_unfold_list_fold_right_cons V W nil_case cons_case v' vs'). rewrite <- IHvs'. rewrite -> (about_list_fold_left_with_left_permutative_cons_case V W cons_case H_left_permutative nil_case v' vs'). reflexivity. Qed. (* {END} *) (* q. Can you think of corollaries of this property? *) (* {task_7_q} *) Lemma plus_is_left_permutative : is_left_permutative nat nat plus. Proof. unfold is_left_permutative. intros v1 v2 w. rewrite -> (Nat.add_assoc v1 v2 w). rewrite -> (Nat.add_assoc v2 v1 w). rewrite -> (Nat.add_comm v2 v1). reflexivity. Qed. Corollary example_for_plus : forall ns : list nat, list_fold_left nat nat 0 plus ns = list_fold_right nat nat 0 plus ns. Proof. Check (folding_left_and_right_over_lists nat nat plus plus_is_left_permutative 0). exact (folding_left_and_right_over_lists nat nat plus plus_is_left_permutative 0). Qed. (* {END} *) (* What do you make of this corollary? Can you think of more such corollaries? *) (* This corollary shows two implementations of a function to sum up lists of nats. Since the order in which addition is performed doesn't matter, it shouldn't matter whether we're folding the list of nats from the left or right. Earlier in this course we learnt that, similarly, the order in which multiplication is performed doesn't matter either i.e. it is commutative. Therefore, we can easily prove it is left permutative and thus prove a corollary of the above list property as well. It is notable that the nil_case should be 1 since mult with 0 as the nil_case is trivial since it always results in 0. *) (* {task_7_q2} *) Lemma mult_is_left_permutative : is_left_permutative nat nat mult. Proof. unfold is_left_permutative. intros v1 v2 w. rewrite -> (Nat.mul_assoc v1 v2 w). rewrite -> (Nat.mul_assoc v2 v1 w). rewrite -> (Nat.mul_comm v2 v1). reflexivity. Qed. Corollary example_for_mult : forall ns : list nat, list_fold_left nat nat 1 mult ns = list_fold_right nat nat 1 mult ns. Proof. Check (folding_left_and_right_over_lists nat nat mult mult_is_left_permutative 1). exact (folding_left_and_right_over_lists nat nat mult mult_is_left_permutative 1). Qed. (* {END} *) (* r. Subsidiary question: does the converse of Theorem folding_left_and_right_over_lists hold? *) (* {task_7_r} *) Theorem folding_left_and_right_over_lists_converse : forall (V W : Type) (cons_case : V -> W -> W), (forall (nil_case : W) (vs : list V), list_fold_left V W nil_case cons_case vs = list_fold_right V W nil_case cons_case vs) -> is_left_permutative V W cons_case. Proof. intros V W cons_case H_fold_eq. unfold is_left_permutative. intros v1 v2 w. rewrite <- (fold_unfold_list_fold_right_nil V W w cons_case) at 1. rewrite <- (fold_unfold_list_fold_right_cons V W w cons_case v2 nil). rewrite <- (fold_unfold_list_fold_right_cons V W w cons_case v1 (v2 :: nil)). rewrite <- (H_fold_eq w (v1 :: v2 :: nil)). rewrite -> (fold_unfold_list_fold_left_cons V W w cons_case v1 (v2 :: nil)). rewrite -> (fold_unfold_list_fold_left_cons V W (cons_case v1 w) cons_case v2 nil). rewrite -> (fold_unfold_list_fold_left_nil V W (cons_case v2 (cons_case v1 w))). reflexivity. Qed. (* {END} *) (* ********** *) (* Task 8: *) (* {nat_fold_right} *) Fixpoint nat_fold_right (V : Type) (zero_case : V) (succ_case : V -> V) (n : nat) : V := match n with O => zero_case | S n' => succ_case (nat_fold_right V zero_case succ_case n') end. (* {END} *) Lemma fold_unfold_nat_fold_right_O : forall (V : Type) (zero_case : V) (succ_case : V -> V), nat_fold_right V zero_case succ_case O = zero_case. Proof. fold_unfold_tactic nat_fold_right. Qed. Lemma fold_unfold_nat_fold_right_S : forall (V : Type) (zero_case : V) (succ_case : V -> V) (n' : nat), nat_fold_right V zero_case succ_case (S n') = succ_case (nat_fold_right V zero_case succ_case n'). Proof. fold_unfold_tactic nat_fold_right. Qed. (* {nat_fold_left} *) Fixpoint nat_fold_left (V : Type) (zero_case : V) (succ_case : V -> V) (n : nat) : V := match n with O => zero_case | S n' => nat_fold_left V (succ_case zero_case) succ_case n' end. (* {END} *) Lemma fold_unfold_nat_fold_left_O : forall (V : Type) (zero_case : V) (succ_case : V -> V), nat_fold_left V zero_case succ_case O = zero_case. Proof. fold_unfold_tactic nat_fold_left. Qed. Lemma fold_unfold_nat_fold_left_S : forall (V : Type) (zero_case : V) (succ_case : V -> V) (n' : nat), nat_fold_left V zero_case succ_case (S n') = nat_fold_left V (succ_case zero_case) succ_case n'. Proof. fold_unfold_tactic nat_fold_left. Qed. (* ***** *) (* The addition function: *) Definition recursive_specification_of_addition (add : nat -> nat -> nat) := (forall y : nat, add O y = y) /\ (forall x' y : nat, add (S x') y = S (add x' y)). Definition tail_recursive_specification_of_addition (add : nat -> nat -> nat) := (forall y : nat, add O y = y) /\ (forall x' y : nat, add (S x') y = add x' (S y)). Definition test_add (candidate: nat -> nat -> nat) : bool := (Nat.eqb (candidate 0 0) 0) && (Nat.eqb (candidate 0 1) 1) && (Nat.eqb (candidate 1 0) 1) && (Nat.eqb (candidate 1 1) 2) && (Nat.eqb (candidate 1 2) 3) && (Nat.eqb (candidate 2 1) 3) && (Nat.eqb (candidate 2 2) 4) && (* commutativity: *) (Nat.eqb (candidate 2 10) (candidate 10 2)) && (* associativity: *) (Nat.eqb (candidate 2 (candidate 5 10)) (candidate (candidate 2 5) 10)) (* etc. *) . (* Testing the unit-test function: *) Compute (test_add Nat.add). Fixpoint r_add (x y : nat) : nat := match x with O => y | S x' => S (r_add x' y) end. Lemma fold_unfold_r_add_O : forall y : nat, r_add O y = y. Proof. fold_unfold_tactic r_add. Qed. Lemma fold_unfold_r_add_S : forall x' y : nat, r_add (S x') y = S (r_add x' y). Proof. fold_unfold_tactic r_add. Qed. (* Implement the addition function as an instance of nat_fold_right or nat_fold_left, your choice. *) Definition r_add_right (x y : nat) : nat := nat_fold_right nat y (fun ih => S ih) x. Compute (test_add r_add_right). Proposition r_add_satisfies_the_recursive_specification_of_addition : recursive_specification_of_addition r_add_right. Proof. unfold recursive_specification_of_addition, r_add_right. split. - intro y. exact (fold_unfold_nat_fold_right_O nat y (fun ih => S ih)). - intros x' y. exact (fold_unfold_nat_fold_right_S nat y (fun ih => S ih) x'). Qed. (* Definition r_add_left (x y : nat) : nat := ... nat_fold_left ... ... ... x ... . *) (* ***** *) (* The power function: *) Definition recursive_specification_of_power (power : nat -> nat -> nat) := (forall x : nat, power x 0 = 1) /\ (forall (x : nat) (n' : nat), power x (S n') = x * power x n'). Definition test_power (candidate : nat -> nat -> nat) : bool := (candidate 2 0 =? 1) && (candidate 10 2 =? 10 * 10) && (candidate 3 2 =? 3 * 3). Fixpoint r_power (x n : nat) : nat := match n with O => 1 | S n' => x * r_power x n' end. Compute (test_power r_power). Lemma fold_unfold_r_power_O : forall x : nat, r_power x O = 1. Proof. fold_unfold_tactic r_power. Qed. Lemma fold_unfold_r_power_S : forall x n' : nat, r_power x (S n') = x * r_power x n'. Proof. fold_unfold_tactic r_power. Qed. Fixpoint tr_power_aux (x n a : nat) : nat := match n with O => a | S n' => tr_power_aux x n' (x * a) end. Lemma fold_unfold_tr_power_aux_O : forall x a : nat, tr_power_aux x 0 a = a. Proof. fold_unfold_tactic tr_power_aux. Qed. Lemma fold_unfold_tr_power_v1_S : forall x n' a : nat, tr_power_aux x (S n') a = tr_power_aux x n' (x * a). Proof. fold_unfold_tactic tr_power_aux. Qed. Definition tr_power (x n : nat) : nat := tr_power_aux x n 1. Compute (test_power tr_power). (* {task_8_b} *) Definition r_power_right (x n : nat) : nat := nat_fold_right nat 1 (fun ih => x * ih) n. Compute (test_power r_power_right). Proposition r_power_right_satisfies_the_recursive_specification_of_power : recursive_specification_of_power r_power_right. Proof. unfold recursive_specification_of_power, r_power_right. split. - intro x. exact (fold_unfold_nat_fold_right_O nat 1 (fun ih : nat => x * ih)). - intros x n'. exact (fold_unfold_nat_fold_right_S nat 1 (fun ih : nat => x * ih) n'). Qed. (* {END} *) (* Definition r_power_left (x n : nat) : nat := ... nat_fold_left ... ... ... n ... . Compute (test_power r_power_left). *) (* Definition tr_power_right (x n : nat) : nat := ... nat_fold_right ... ... ... n ... . Compute (test_power tr_power_right). *) (* Definition tr_power_left (x n : nat) : nat := ... nat_fold_left ... ... ... n ... . Compute (test_power tr_power_left). *) (* ***** *) (* The factorial function: *) Definition recursive_specification_of_the_factorial_function (fac : nat -> nat) := (fac 0 = 1) /\ (forall n' : nat, fac (S n') = S n' * fac n'). Definition test_fac (candidate : nat -> nat) : bool := (candidate 0 =? 1) && (candidate 1 =? 1 * 1) && (candidate 2 =? 2 * 1 * 1) && (candidate 3 =? 3 * 2 * 1 * 1) && (candidate 4 =? 4 * 3 * 2 * 1 * 1) && (candidate 5 =? 5 * 4 * 3 * 2 * 1 * 1). Fixpoint r_fac (n : nat) : nat := match n with O => 1 | S n' => S n' * r_fac n' end. Compute (test_fac r_fac). Lemma fold_unfold_r_fac_O : r_fac 0 = 1. Proof. fold_unfold_tactic r_fac. Qed. Lemma fold_unfold_r_fac_S : forall n' : nat, r_fac (S n') = S n' * r_fac n'. Proof. fold_unfold_tactic r_fac. Qed. Proposition r_fac_satisfies_the_recursive_specification_of_the_factorial_function : recursive_specification_of_the_factorial_function r_fac. Proof. unfold recursive_specification_of_the_factorial_function. exact (conj fold_unfold_r_fac_O fold_unfold_r_fac_S). Qed. (* Re-implement r_fac as an instance of nat_fold_right or nat_fold_left, your choice: *) (* {task_8_c} *) Definition r_fac_right (n : nat) : nat := fst (nat_fold_right (nat * nat) (1, 0) (fun p => (fst p * S (snd p), S (snd p))) n). Compute (test_fac r_fac_right). Lemma about_r_fac_right : forall (n : nat), snd (nat_fold_right (nat * nat) (1, 0) (fun p => (fst p * S (snd p), S (snd p))) n) = n. Proof. intro n. remember (fun p : nat * nat => (fst p * S (snd p), S (snd p))) as f eqn:H_f. induction n as [ | n' IHn']. - rewrite -> (fold_unfold_nat_fold_right_O (nat * nat) (1, 0) f). unfold snd. reflexivity. - rewrite -> (fold_unfold_nat_fold_right_S (nat * nat) (1, 0) f n'). rewrite -> H_f. rewrite <- H_f. unfold snd at 1. rewrite -> IHn'. reflexivity. Qed. Proposition r_fac_right_satisfies_the_recursive_specification_of_the_factorial_function : recursive_specification_of_the_factorial_function r_fac_right. Proof. unfold recursive_specification_of_the_factorial_function, r_fac_right. split. - remember (fun p : nat * nat => (fst p * S (snd p), S (snd p))) as f eqn:H_f. rewrite -> (fold_unfold_nat_fold_right_O (nat * nat) (1, 0) f). unfold fst. reflexivity. - remember (fun p : nat * nat => (fst p * S (snd p), S (snd p))) as f eqn:H_f. intro n'. rewrite -> (fold_unfold_nat_fold_right_S (nat * nat) (1, 0) f n'). rewrite -> H_f. rewrite <- H_f. unfold fst at 1. rewrite -> H_f at 2. rewrite -> (about_r_fac_right n'). exact (Nat.mul_comm (fst (nat_fold_right (nat * nat) (1, 0) f n')) (S n')). Qed. (* {END} *) (* Definition fac_left (n : nat) : nat := ... nat_fold_left ... ... ... n ... . Compute (test_fac r_fac_left). *) Fixpoint tr_fac_aux (n a : nat) : nat := match n with O => a | S n' => tr_fac_aux n' (S n' * a) end. Definition tr_fac (n : nat) : nat := tr_fac_aux n 1. Compute (test_fac tr_fac). Lemma fold_unfold_tr_fac_aux_O : forall a : nat, tr_fac_aux 0 a = a. Proof. fold_unfold_tactic tr_fac_aux. Qed. Lemma fold_unfold_tr_fac_aux_S : forall n' a : nat, tr_fac_aux (S n') a = tr_fac_aux n' (S n' * a). Proof. fold_unfold_tactic tr_fac_aux. Qed. (* Re-implement tr_fac as an instance of nat_fold_right or nat_fold_left, your choice: *) (* {task_8_c2} *) Definition tr_fac_left (n : nat) := fst (nat_fold_left (nat * nat) (1, 0) (fun p => (fst p * S (snd p), S (snd p))) n). Compute (test_fac tr_fac_left). Proposition about_tr_fac_left : forall (n x ih : nat), fst (nat_fold_left (nat * nat) (S x * ih, S x) (fun p : nat * nat => (fst p * S (snd p), S (snd p))) n) = S (n + x) * fst (nat_fold_left (nat * nat) (ih, x) (fun p : nat * nat => (fst p * S (snd p), S (snd p))) n). Proof. intros n. remember (fun p : nat * nat => (fst p * S (snd p), S (snd p))) as f eqn:H_f. induction n as [ | n' IHn']. - intros x ih. rewrite -> (fold_unfold_nat_fold_left_O (nat * nat) (ih, x) f). rewrite -> (fold_unfold_nat_fold_left_O (nat * nat) (S x * ih, S x) f). unfold fst. rewrite -> (Nat.add_0_l x). reflexivity. - intros x ih. rewrite -> (fold_unfold_nat_fold_left_S (nat * nat) (S x * ih, S x) f n'). rewrite -> (fold_unfold_nat_fold_left_S (nat * nat) (ih, x) f n'). rewrite -> H_f. rewrite <- H_f. unfold snd. unfold fst at 2 4. rewrite -> (Nat.mul_shuffle0 (S x) ih (S (S x))). rewrite -> (Nat.mul_comm (S x) (S (S x))). rewrite <- (Nat.mul_assoc (S (S x)) (S x) ih). rewrite -> (IHn' (S x) (S x * ih)). rewrite -> (Nat.add_succ_comm n' x). rewrite -> (Nat.mul_comm ih (S x)). reflexivity. Qed. Proposition tr_fac_left_satisfies_the_recursive_specification_of_the_factorial_function : recursive_specification_of_the_factorial_function tr_fac_left. Proof. unfold recursive_specification_of_the_factorial_function, tr_fac_left. split. - remember (fun p : nat * nat => (fst p * S (snd p), S (snd p))) as f eqn:H_f. rewrite -> (fold_unfold_nat_fold_left_O (nat * nat) (1, 0) f). unfold fst. reflexivity. - remember (fun p : nat * nat => (fst p * S (snd p), S (snd p))) as f eqn:H_f. intro n'. rewrite -> (fold_unfold_nat_fold_left_S (nat * nat) (1, 0) f n'). rewrite -> H_f. rewrite <- H_f. unfold fst at 2. unfold snd. rewrite -> H_f. Check (about_tr_fac_left (n') 0 1). rewrite <- (Nat.add_0_r n') at 2. exact (about_tr_fac_left (n') 0 1). Qed. (* {END} *) (* Definition tr_fac_right (n : nat) : nat := ... nat_fold_right ... ... ... n ... . Compute (test_fac tr_fac_right). Definition tr_fac_left (n : nat) : nat := ... nat_fold_left ... ... ... n ... . Compute (test_fac tr_fac_alt). *) (* ***** *) Definition specification_of_the_fibonacci_function (fib : nat -> nat) := fib 0 = 0 /\ fib 1 = 1 /\ forall n'' : nat, fib (S (S n'')) = fib (S n'') + fib n''. Definition test_fib (candidate: nat -> nat) : bool := (candidate 0 =? 0) && (candidate 1 =? 1) && (candidate 2 =? 1) && (candidate 3 =? 2) && (candidate 4 =? 3) && (candidate 5 =? 5) (* etc. *). Fixpoint r_fib (n : nat) : nat := match n with 0 => 0 | S n' => match n' with 0 => 1 | S n'' => r_fib n' + r_fib n'' end end. Compute (test_fib r_fib). Lemma fold_unfold_r_fib_O : r_fib O = 0. Proof. fold_unfold_tactic r_fib. Qed. Lemma fold_unfold_r_fib_S : forall n' : nat, r_fib (S n') = match n' with 0 => 1 | S n'' => r_fib n' + r_fib n'' end. Proof. fold_unfold_tactic r_fib. Qed. Corollary fold_unfold_r_fib_SO : r_fib 1 = 1. Proof. rewrite -> (fold_unfold_r_fib_S 0). reflexivity. Qed. Corollary fold_unfold_r_fib_SS : forall n'' : nat, r_fib (S (S n'')) = r_fib (S n'') + r_fib n''. Proof. intro n''. rewrite -> (fold_unfold_r_fib_S (S n'')). reflexivity. Qed. Proposition r_fib_satisfies_the_specification_of_the_fibonacci_function : specification_of_the_fibonacci_function r_fib. Proof. unfold specification_of_the_fibonacci_function. exact (conj fold_unfold_r_fib_O (conj fold_unfold_r_fib_SO fold_unfold_r_fib_SS)). Qed. (* Implement the Fibonacci function as an instance of nat_fold_right or nat_fold_left, your choice: *) (* {task_8_d} *) Definition fib_right_succ (p : nat * nat) : (nat * nat) := (snd p, fst p + snd p). Definition fib_right (n : nat) : nat := fst (nat_fold_right (nat * nat) (0, 1) fib_right_succ n). Compute (test_fib fib_right). Proposition fib_right_satisfies_the_specification_of_the_fibonacci_function : specification_of_the_fibonacci_function fib_right. Proof. unfold specification_of_the_fibonacci_function, fib_right. split. - rewrite -> (fold_unfold_nat_fold_right_O (nat * nat) (0, 1) fib_right_succ). unfold fst. reflexivity. - split. + rewrite -> (fold_unfold_nat_fold_right_S (nat * nat) (0, 1) fib_right_succ 0). unfold fib_right_succ at 1. unfold fst at 1. rewrite -> (fold_unfold_nat_fold_right_O (nat * nat) (0, 1) fib_right_succ). unfold snd. reflexivity. + intro n''. rewrite -> (fold_unfold_nat_fold_right_S (nat * nat) (0, 1) fib_right_succ (S n'')). unfold fib_right_succ at 1. unfold fst at 1. rewrite -> (fold_unfold_nat_fold_right_S (nat * nat) (0, 1) fib_right_succ n''). unfold fib_right_succ at 1. unfold snd at 1. unfold fib_right_succ at 3. unfold fst at 2. exact (Nat.add_comm (fst (nat_fold_right (nat * nat) (0, 1) fib_right_succ n'')) (snd (nat_fold_right (nat * nat) (0, 1) fib_right_succ n''))). Qed. (* {END} *) (* Definition fib_left (n : nat) : nat := ... nat_fold_left ... ... ... n ... . Compute (test_fib fib_left). *) (* ********** *) (* Task 9 *) (* Under which conditions -- if any -- are nat_fold_right and nat_fold_left equivalent? *) (* {task_9} *) Lemma about_nat_fold_right : forall (V : Type) (zero_case : V) (succ_case : V -> V) (n : nat), nat_fold_right V (succ_case zero_case) succ_case n = succ_case (nat_fold_right V zero_case succ_case n). Proof. intros V zero_case succ_case n. induction n as [ | n' IHn']. - rewrite -> (fold_unfold_nat_fold_right_O V (succ_case zero_case) succ_case). rewrite -> (fold_unfold_nat_fold_right_O V zero_case succ_case). reflexivity. - rewrite -> (fold_unfold_nat_fold_right_S V (succ_case zero_case) succ_case n'). rewrite -> (fold_unfold_nat_fold_right_S V zero_case succ_case n'). rewrite -> IHn'. reflexivity. Qed. Proposition folding_left_and_right_over_nats : forall (V : Type) (zero_case : V) (succ_case : V -> V) (n : nat), nat_fold_left V zero_case succ_case n = nat_fold_right V zero_case succ_case n. Proof. intros V zero_case succ_case n. revert zero_case. induction n as [ | n' IHn']. - intro zero_case. rewrite -> (fold_unfold_nat_fold_left_O V zero_case succ_case). rewrite -> (fold_unfold_nat_fold_right_O V zero_case succ_case). reflexivity. - intro zero_case. rewrite -> (fold_unfold_nat_fold_left_S V zero_case succ_case n'). rewrite -> (fold_unfold_nat_fold_right_S V zero_case succ_case n'). rewrite -> (IHn' (succ_case zero_case)). rewrite -> (about_nat_fold_right V zero_case succ_case n'). reflexivity. Qed. (* {END} *) (* ********** *) (* end of midterm_project.v *) (* It does this by removing the first element from vs and cons-ing it with the accumulator and then calling itself with the vs' and the longer accumulator *) Fixpoint list_reverse_acc_1 (V : Type) (vs a : list V) : list V := match vs with nil => a | v :: vs' => list_reverse_acc V vs' (v :: a) end. (* This states that the reverse of 2 lists, v1s and v2s that are appended is equivalent to the reversed v2s appended to the reversed v1s *) (* Proposition list_reverse_acc_and_list_append_commute_with_each_other_1 : *) (* forall (V : Type) *) (* (v1s v2s: list V), *) (* list_append V (list_reverse_acc V v2s nil) (list_reverse_acc V v1s nil) = *) (* list_reverse_acc V (list_append V v1s v2s) nil. *) (* Admitted. *) (* This generalises the above by stating that *) (* RHS: append( (append reversed v2s reversed v1s) v3s) *) (* LHS: append v1s and v2s, reverse that entire thing, which should give *) (* append (reversed v2s) (reversed v1s)*) (* This is then appended *) Proposition a_generalized_alternative : forall (V : Type) (v1s v2s v3s : list V), list_reverse_acc V (list_append V v1s v2s) v3s = (list_reverse_acc V v2s (list_reverse_acc V v1s v3s)). Proof. Compute (let V := nat in let v1s := 11 :: 12 :: nil in let v2s := 21 :: 22 :: nil in let v3s := 31 :: 32 :: nil in list_reverse_acc V (list_append V v1s v2s) v3s = list_reverse_acc V v2s (list_reverse_acc V v1s v3s) ). (* 22 :: 21 :: 12 :: 11 :: 31 :: 32 :: nil *) Compute (let V := nat in let v1s := 11 :: 12 :: nil in let v2s := 21 :: 22 :: nil in let v3s := 31 :: 32 :: nil in eqb_list nat Nat.eqb (list_reverse_acc V (list_append V v1s v2s) v3s) (list_reverse_acc V v2s (list_reverse_acc V v1s v3s))). intros V v1s. induction v1s as [| v1 v1s' IHv1s']. - intros v2s v3s. rewrite -> (fold_unfold_list_append_nil V v2s). rewrite -> (fold_unfold_list_reverse_acc_nil V v3s). reflexivity. - intros v2s v3s. (* reverse_acc (11::12::21::22::nil) (31::32::nil) ) = *) (* reverse_acc (21::22::nil) (reverse_acc (11::12::nil) (31::32::nil) *) Search list_reverse_acc. rewrite -> (fold_unfold_list_append_cons V v1 v1s'). rewrite -> (fold_unfold_list_reverse_acc_cons V v1 (list_append V v1s' v2s)). rewrite -> (fold_unfold_list_reverse_acc_cons V v1 v1s'). exact (IHv1s' v2s (v1 :: v3s)). Restart. intros V v1s v2s. induction v1s as [| v1 v1s' IHv1s']. - intro v3s. rewrite -> (fold_unfold_list_append_nil V v2s). rewrite -> (fold_unfold_list_reverse_acc_nil V v3s). reflexivity. - intros v3s. (* reverse_acc (11::12::21::22::nil) (31::32::nil) ) = *) (* reverse_acc (21::22::nil) (reverse_acc (11::12::nil) (31::32::nil) *) Search list_reverse_acc. rewrite -> (fold_unfold_list_append_cons V v1 v1s'). rewrite -> (fold_unfold_list_reverse_acc_cons V v1 (list_append V v1s' v2s)). rewrite -> (fold_unfold_list_reverse_acc_cons V v1 v1s'). exact (IHv1s' (v1 :: v3s)). Restart. intros V v1s v2s v3s. Search list_reverse_acc. rewrite -> (about_list_reverse_acc_and_append V (list_append V v1s v2s) v3s). Search list_reverse_acc. rewrite <- (list_reverse_acc_and_list_append_commute_with_each_other V v1s v2s). Search (list_append _ (list_append _ _ _) _). rewrite <- list_append_is_associative. Search list_reverse_acc. Check about_list_reverse_acc_and_append. rewrite <- (about_list_reverse_acc_and_append V v1s v3s). rewrite <- (about_list_reverse_acc_and_append V v2s (list_reverse_acc V v1s v3s)). reflexivity. Qed. Proposition a_generalization : forall (V : Type) (v1s v2s v3s : list V), list_reverse_acc V (list_reverse_acc V v1s v2s) v3s = list_reverse_acc V v2s (list_append V v1s v3s). Proof. Compute (let V := nat in let v1s := 11 :: 12 :: nil in let v2s := 21 :: 22 :: nil in let v3s := 31 :: 32 :: nil in list_reverse_acc V (list_reverse_acc V v1s v2s) v3s = list_reverse_acc V v2s (list_append V v1s v3s)). Compute (let V := nat in let v1s := 11 :: 12 :: nil in let v2s := 21 :: 22 :: nil in let v3s := 31 :: 32 :: nil in eqb_list nat Nat.eqb (list_reverse_acc V (list_reverse_acc V v1s v2s) v3s) (list_reverse_acc V v2s (list_append V v1s v3s))). intros V v1s v2s v3s. revert v2s. induction v1s as [ | v1 v1s' IHv1s' ]. - intros v2s. rewrite -> (fold_unfold_list_reverse_acc_nil V v2s). rewrite -> (fold_unfold_list_append_nil V v3s). reflexivity. - intros v2s. rewrite -> (fold_unfold_list_append_cons). rewrite -> (fold_unfold_list_reverse_acc_cons). rewrite -> (IHv1s' (v1 :: v2s)). rewrite -> (fold_unfold_list_reverse_acc_cons). reflexivity. Restart. intros V v1s v2s v3s. rewrite -> (about_list_reverse_acc_and_append V v2s (list_append V v1s v3s)). Check list_append_is_associative. rewrite -> list_append_is_associative. rewrite <- (about_list_reverse_acc_and_append V v2s v1s). rewrite -> (about_list_reverse_acc_and_append V (list_reverse_acc V v1s v2s) v3s). Search (list_reverse_acc _ (list_reverse_acc _ _ _) _). assert (helper : (list_reverse_acc V (list_reverse_acc V v1s v2s) nil) = (list_reverse_acc V v2s v1s)). { rewrite -> (about_list_reverse_acc_and_append V v1s v2s). Search list_reverse_acc. rewrite <- list_reverse_acc_and_list_append_commute_with_each_other. rewrite -> list_reverse_acc_is_involutory. rewrite <- about_list_reverse_acc_and_append. reflexivity. } rewrite -> helper. reflexivity. Qed.