diff --git a/cs3234/labs/midterm-project.v b/cs3234/labs/midterm-project.v new file mode 100644 index 0000000..f154988 --- /dev/null +++ b/cs3234/labs/midterm-project.v @@ -0,0 +1,1415 @@ +(* 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: + date: + + 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 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: *) + +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. + intros S_eqb_V. + intros v1s. + induction v1s as [ | v1 v1s' IHv1s' ]. + - intros v2s H_eqb_list. + case v2s as [ | v2 v2s' ]. + + reflexivity. + + rewrite -> fold_unfold_eqb_list_nil in H_eqb_list. + discriminate H_eqb_list. + - intros v2s H_eqb_list. + case v2s as [ | v2 v2s' ]. + + rewrite -> fold_unfold_eqb_list_cons in H_eqb_list. + discriminate H_eqb_list. + + rewrite -> fold_unfold_eqb_list_cons in H_eqb_list. + Search (_ && _ = true). + destruct (andb_true_iff (eqb_V v1 v2) (eqb_list V eqb_V v1s' v2s')) as [H_tmp _]. + destruct (H_tmp H_eqb_list) as [v1_eq_v2 v1s'_eq_v2s']. + rewrite -> (S_eqb_V v1 v2 v1_eq_v2). + rewrite -> (IHv1s' v2s' v1s'_eq_v2s'). + 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 v2s H_eqb_list. + case v2s as [ | v2 v2s' ]. + + rewrite -> fold_unfold_eqb_list_nil. + reflexivity. + + discriminate H_eqb_list. + - intros v2s v1s_eq_v2s. + case v2s as [ | v2 v2s' ]. + + discriminate v1s_eq_v2s. + + rewrite -> fold_unfold_eqb_list_cons. + injection v1s_eq_v2s as v1_eq_v2 v1s'_eq_v2s'. + Check (C_eqb_V v1 v2 v1_eq_v2). + rewrite -> (C_eqb_V v1 v2 v1_eq_v2). + rewrite -> (IHv1s' v2s' v1s'_eq_v2s'). + simpl (true && true). + reflexivity. +Qed. + +(* ********** *) + +(* A study of the polymorphic length function: *) + +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')). + +(* 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: *) + +(* Implement the length function using an accumulator. *) + +Fixpoint list_length_acc (V : Type) (vs : list V) (acc : nat) : nat := + match vs with + nil => acc + | v :: vs' => + list_length_acc V vs' (S acc) + end. + +Definition list_length_alt (V : Type) (vs : list V) : nat := + list_length_acc V vs 0. + +Compute (test_list_length list_length_alt). + +Lemma fold_unfold_list_length_acc_nil : + forall (V : Type) + (acc : nat), + list_length_acc V nil acc = acc. +Proof. + fold_unfold_tactic list_length_acc. +Qed. + +Lemma fold_unfold_list_length_acc_cons : + forall (V : Type) + (v : V) + (vs' : list V) + (acc : nat), + list_length_acc V (v :: vs') acc = list_length_acc V vs' (S acc). +Proof. + fold_unfold_tactic list_length_acc. +Qed. + +Lemma about_list_length_acc : + forall (V : Type) + (vs : list V) + (acc : nat), + list_length_acc V vs (S acc) = S (list_length_acc V vs acc). +Proof. + intros V vs. + induction vs as [ | v vs' IHvs' ]. + - intro acc. + rewrite ->(fold_unfold_list_length_acc_nil V (S acc)). + rewrite -> (fold_unfold_list_length_acc_nil V acc). + reflexivity. + - intro acc. + rewrite -> (fold_unfold_list_length_acc_cons V v vs' (S acc)). + rewrite -> (IHvs' (S acc)). + rewrite -> (fold_unfold_list_length_acc_cons V v vs' acc). + reflexivity. +Qed. + +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. + - intros 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' 0). + reflexivity. +Qed. + +(* ********** *) + +(* A study of the polymorphic copy function: *) + +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')). + +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. +*) + +(* + b. Implement the copy function recursively. +*) + +(* +Fixpoint list_copy (V : Type) (vs : list V) : list V := + ... +*) + +(* + c. State its associated fold-unfold lemmas. +*) + +(* + d. Prove whether your implementation satisfies the specification. +*) + +(* + e. Prove whether copy is idempotent. +*) + +(* +Proposition list_copy_is_idempotent : + forall (V : Type) + (vs : list V), + list_copy V (list_copy V vs) = list_copy V vs. +Proof. +Abort. +*) + +(* + f. Prove whether copying a list preserves its length. +*) + +(* +Proposition list_copy_preserves_list_length : + forall (V : Type) + (vs : list V), + list_length V (list_copy V vs) = list_length V vs. +Proof. +Abort. +*) + +(* + 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? +*) + +(* ********** *) + +(* 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. +*) + +(* + b. Implement the list_append function recursively. +*) + +(* +Fixpoint list_append (V : Type) (v1s v2s : list V) : list V := + ... +*) + +(* + c. State its associated fold-unfold lemmas. +*) + +(* + d. Prove that your implementation satisfies the specification. +*) + +(* + e. Prove whether nil is neutral on the left of list_append. +*) + +(* + f. Prove whether nil is neutral on the right of list_append. +*) + +(* + g. Prove whether list_append is commutative. +*) + +(* + h. Prove whether list_append is associative. +*) + +(* + i. Prove whether appending two lists preserves their length. +*) + +(* +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. +Abort. +*) + +(* + j. Prove whether list_append and list_copy commute with each other. +*) + +(* +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. +Abort. +*) + +(* ********** *) + +(* A study of the polymorphic reverse function: *) + +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)). + +(* ***** *) + +(* Task 5: *) + +(* + a. Define a unit-test function for an implementation of the reverse function. +*) + +(* + b. Implement the reverse function recursively, using list_append. +*) + +(* +Fixpoint list_reverse (V : Type) (vs : list V) : list V := +... +*) + +(* + c. State the associated fold-unfold lemmas. +*) + +(* + d. Prove whether your implementation satisfies the specification. +*) + +(* + e. Prove whether list_reverse is involutory. +*) + +(* +Proposition list_reverse_is_involutory : + forall (V : Type) + (vs : list V), + list_reverse V (list_reverse V vs) = vs. +Proof. +Abort. +*) + +(* + f. Prove whether reversing a list preserves its length. +*) + +(* + g. Do list_append and list_reverse commute with each other (hint: yes they do) and if so how? +*) + +(* + h. Implement the reverse function using an accumulator instead of using list_append. +*) + +(* +Definition list_reverse_alt (V : Type) (vs : list V) : list V := +... +*) + +(* + 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. +*) + +(* ********** *) + +(* A study of the polymorphic map function: *) + +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'). + +(* ***** *) + +(* Task 6: + + a. Prove whether the specification specifies at most one map function. +*) + +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. + +(* + b. Implement the map function recursively. +*) + +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. +*) + +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. + +(* + d. Prove whether your implementation satisfies the specification. +*) + +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. + +(* + e. Implement the copy function as an instance of list_map. +*) + +(* +Definition list_copy_as_list_map (V : Type) (vs : list V) : list V := + ... +*) + +(* +Hint: Does list_copy_as_list_map satisfy the specification of list_copy? +*) + +(* + f. Prove whether mapping a function over a list preserves the length of this list. +*) + +(* + g. Do list_map and list_append commute with each other and if so how? +*) + +(* + h. Do list_map and list_reverse commute with each other and if so how? +*) + +(* + i. Do list_map and list_reverse_alt commute with each other and if so how? +*) + +(* + j. Define a unit-test function for the map function + and verify that your implementation satisfies it. +*) + +(* ********** *) + +(* A study of the polymorphic fold-right and fold-left functions: *) + +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'). + +(* ***** *) + +(* Task 7: + + a. Implement the fold-right function recursively. +*) + +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. + +(* + b. Implement the fold-left function recursively. +*) + +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. + +(* + c. state the fold-unfold lemmas associated to list_fold_right and to list_fold_left +*) + +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. + +(* + d. Prove that each of your implementations satisfies the corresponding specification. +*) + +(* + e. Which function do foo and bar (defined just below) compute? +*) + +(* +Definition foo (V : Type) (vs : list V) := + list_fold_right V (list V) nil (fun v vs => v :: vs) vs. +*) + +(* +Definition bar (V : Type) (vs : list V) := + list_fold_left V (list V) nil (fun v vs => v :: vs) vs. +*) + +(* + 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. +*) + +(* + 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. +*) + +(* + 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. +*) + +(* + 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. +*) + +(* + 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. +*) + +(* + k. Implement eqb_list either as an instance of list_fold_right or as an instance of list_fold_left, and justify your choice. +*) + +(* + l. Implement list_fold_right as an instance of list_fold_left, using list_reverse. +*) + +(* + m. Implement list_fold_left as an instance of list_fold_right, using list_reverse. +*) + +(* + n. Implement list_fold_right as an instance of list_fold_left, without using list_reverse. +*) + +(* +Definition list_fold_right_left (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (vs : list V) : W := + ... +*) + +(* + o. Implement list_fold_left as an instance of list_fold_right, without using list_reverse. +*) + +(* +Definition list_fold_left_right (V W : Type) (nil_case : W) (cons_case : V -> W -> W) (vs : list V) : W := + ... +*) + +(* + 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 +*) + +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). + +(* +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. +Abort. +*) + +(* + q. Can you think of corollaries of this property? +*) + +Lemma plus_is_left_permutative : + is_left_permutative nat nat plus. +Proof. +Abort. + +(* +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. +*) + +(* What do you make of this corollary? + Can you think of more such corollaries? +*) + +(* + r. Subsidiary question: does the converse of Theorem folding_left_and_right_over_lists hold? +*) + +(* +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. +Abort. +*) + +(* ********** *) + +(* Task 8: *) + +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. + +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. + +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. + +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). + +(* +Definition r_power_right (x n : nat) : nat := + ... nat_fold_right ... ... ... n ... . + +Compute (test_power r_power_right). +*) + +(* +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: *) + +(* +Definition r_fac_right (n : nat) : nat := + ... nat_fold_right ... ... ... n ... . + +Compute (test_fac r_fac_right). + +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. + +Proposition tr_fac_satisfies_the_recursive_specification_of_the_factorial_function : + recursive_specification_of_the_factorial_function tr_fac. +Proof. + unfold recursive_specification_of_the_factorial_function. +Abort. + +(* Re-implement tr_fac as an instance of nat_fold_right or nat_fold_left, your choice: *) + +(* +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: *) + +(* +Definition fib_right (n : nat) : nat := + ... nat_fold_right ... ... ... n ... . + +Compute (test_fib tr_fib_right). + +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? *) + +(* ********** *) + +(* end of midterm-project.v *) diff --git a/cs3234/labs/midterm_project_copy-for-oral-exam.v b/cs3234/labs/midterm_project_copy-for-oral-exam.v new file mode 100644 index 0000000..7d946ad --- /dev/null +++ b/cs3234/labs/midterm_project_copy-for-oral-exam.v @@ -0,0 +1,3657 @@ +(* 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. diff --git a/cs3234/labs/question-for-Yadunand-Prem-midterm.txt b/cs3234/labs/question-for-Yadunand-Prem-midterm.txt new file mode 100644 index 0000000..a676d7b --- /dev/null +++ b/cs3234/labs/question-for-Yadunand-Prem-midterm.txt @@ -0,0 +1,85 @@ +% question-for-Yadunand-Prem.txt +% CS3234, Fri 05 Apr 2024 + +%%%%%%%%%% + +This oral exam is a continuation of the midterm project. +So, make a copy of midterm_project.v (naming it midterm_project_copy-for-oral-exam.v) +and start editing this file after + +(* end of midterm_project.v *) + +%%%%%%%%%% + +The goal of this oral exam is to revisit + + 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. + +So, copy the following proposition at the very end of midterm_project_copy-for-oral-exam.v: + + Proposition a_generalized_alternative : + forall (V : Type) + (v1s v2s v3s : list V), + list_reverse_acc V (list_append V v1s v2s) v3s + = nil. + +%%%%% + +a. The first thing to do is to figure out what to put instead of nil on the right-hand side, + so that the equality holds. + We (in the sense of you) are going to do that empirically by testing. + To this end, write: + + 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 + = nil). + + Then, replace nil on the right-hand-side by an expression that does not use list_append + but that only uses two occurrences of list_reverse_acc, + and that is such that the equality holds. + (Let us name this expression as YOUR_EXPRESSION.) + + Hint: use "(list_reverse_acc V ... ...)" as an accumulator for list_reverse_acc. + +b. To make sure, write: + + 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) + YOUR_EXPRESSION). + + Verify that the result of this computation is true. + If it is false, then go back to a. and revise YOUR_EXPRESSION + until the result of the computation just above is true. + +c. Prove this equality by induction. + +d. Time permitting, write: + + Restart. + + Prove these equalities not by induction, + but using the many properties you have proved earlier in the file, + e.g., about_list_reverse_acc_and_append, + list_reverse_acc_and_list_append_commute_with_each_other, + list_append_is_associative, + etc. + (Maybe you will need to state and prove more similar properties to carry our this proof, + maybe not, who knows, this is an exploration.) + +%%%%%%%%%% + +% end of question-for-Yadunand-Prem.txt diff --git a/cs3234/labs/week-06_a_simple_problem.v b/cs3234/labs/week-06_a_simple_problem.v new file mode 100644 index 0000000..d96a109 --- /dev/null +++ b/cs3234/labs/week-06_a_simple_problem.v @@ -0,0 +1,72 @@ +(* Paraphernalia: *) + +Ltac fold_unfold_tactic name := intros; unfold name; fold name; reflexivity. + +Require Import Arith Bool. + +(* ********** *) + +Check Bool.eqb. (* : bool -> bool -> bool *) + +Check eqb. (* : bool -> bool -> bool *) + +Search (eqb _ _ = true -> _ = _). +(* eqb_prop: forall a b : bool, eqb a b = true -> a = b *) + +Search (eqb _ _ = true). +(* eqb_reflx: forall b : bool, eqb b b = true *) + +Search (_ * _). + +(* Definition is_even (n: nat) := *) + +(* Proposition product_of_n_and_even_is_even : *) +(* forall n : nat, *) + +Proposition product_of_2_consecutive_natural_numbers_is_even: + forall n: nat, + exists a: nat, + n * (S n) = 2 * a. +Proof. + intro n. + induction n as [ | n' IHn']. + - exists 0. + rewrite -> (Nat.mul_0_l 1). + rewrite -> (Nat.mul_0_r 2). + reflexivity. + - rewrite -> (Nat.mul_succ_r (S n') (S n')). + rewrite -> (Nat.mul_succ_r (S n') n'). + rewrite -> (Nat.mul_comm (S n') n'). + destruct IHn' as [k IHn']. + rewrite -> IHn'. + rewrite <- (Nat.add_assoc (2 * k) (S n') (S n')). + remember (S n' + S n') as x eqn:H_x. + rewrite <- (Nat.mul_1_l (S n')) in H_x at 1. + rewrite -> H_x. + rewrite <- (Nat.mul_succ_l 1 (S n')). + Search (_ * _ + _ * _). + rewrite <- (Nat.mul_add_distr_l 2 k (S n')). + exists (k + S n'). + reflexivity. +Qed. + +Proposition product_of_3_consecutive_natural_numbers_is_divisible_by_2: + forall n : nat, + exists a : nat, + n * (S n) * (S (S n)) = 2 * a. +Proof. + intro n. + induction n as [ | n' IHn' ]. + - exists 0. + rewrite -> (Nat.mul_0_l 1). + rewrite -> (Nat.mul_0_l 2). + rewrite -> (Nat.mul_0_r 2). + reflexivity. + - destruct IHn' as [k IHn']. + destruct (product_of_2_consecutive_natural_numbers_is_even (S n')) as [x H_product_of_2_consecutive_natural_numbers_is_even]. + rewrite -> H_product_of_2_consecutive_natural_numbers_is_even. + exists (x * S (S (S n'))). + Check (Nat.mul_assoc 2 x (S (S (S n')))). + rewrite -> (Nat.mul_assoc 2 x (S (S (S n')))). + reflexivity. +Qed. diff --git a/cs3234/labs/week-06_ex-falso-quodlibet.v b/cs3234/labs/week-06_ex-falso-quodlibet.v new file mode 100644 index 0000000..84d1494 --- /dev/null +++ b/cs3234/labs/week-06_ex-falso-quodlibet.v @@ -0,0 +1,172 @@ +(* week-06_ex-falso-quodlibet.v *) +(* LPP 2024 - CS3234 2023-2024, Sem2 *) +(* Olivier Danvy *) +(* Version of 22 Feb 2024 *) + +(* ********** *) + +Require Import Arith. + +(* ********** *) + +Property foo : + forall P : nat -> Prop, + (exists i : nat, + P i) -> + forall j : nat, + P j. +Proof. +Abort. (* does not hold, see just below *) + +Theorem ex_falso_quodlibet_eg_False : + (forall P : nat -> Prop, + (exists i : nat, + P i) -> + forall j : nat, + P j) -> + 0 = 1. +Proof. + intro H_absurd. + Check (H_absurd + (fun n : nat => 0 = n)). + (* : (exists i : nat, 0 = i) -> forall j : nat, 0 = j *) + (* wanted: exists i : nat, 0 = i *) + Check ex_intro. + Check (ex_intro + (fun i : nat => 0 = i)). + (* : forall x : nat, 0 = x -> exists i : nat, 0 = i *) + (* let's pick 0, for example *) + Check (ex_intro + (fun i : nat => 0 = i) + 0). + Check (ex_intro + (fun i : nat => 0 = i) + 0 + (eq_refl 0)). + (* : exists i : nat, 0 = i *) + (* which is what we wanted *) + Check (H_absurd + (fun n : nat => 0 = n) + (ex_intro + (fun i : nat => 0 = i) + 0 + (eq_refl 0))). + (* : forall j : nat, 0 = j *) + (* let's pick 1, for example *) + Check (H_absurd + (fun n : nat => 0 = n) + (ex_intro + (fun i : nat => 0 = i) + 0 + (eq_refl 0)) + 1). + (* : 0 = 1 *) + exact (H_absurd + (fun n : nat => 0 = n) + (ex_intro + (fun i : nat => 0 = i) + 0 + (eq_refl 0)) + 1). +Qed. + +(* ********** *) + +(* Exercise 06 *) + +Proposition ex_falso_quodlibet_indeed : + (forall (P : nat -> Prop), + (exists i : nat, + P i) -> + forall j : nat, + P j) -> False. +Proof. +Abort. + +(* ********** *) + +(* Exercise 07 *) + +Property foo : (* stated again, but admitted this time... *) + forall P : nat -> Prop, + (exists i : nat, + P i) -> + forall j : nat, + P j. +Proof. +Admitted. (* ...to prove the following theorem as a corollary of foo *) + +Theorem ex_falso_quodlibet_eg_True : + forall m : nat, + exists n : nat, + m < n. +Proof. + Check (foo (fun m : nat => exists n : nat, m < n)). +Abort. + +(* ********** *) + +(* Exercise 08 *) + +Proposition an_equivalence : + (forall P : nat -> Prop, + (exists i : nat, + P i) -> + forall j : nat, + P j) + <-> + (forall (P : nat -> Prop) + (i : nat), + P i -> + forall j : nat, + P j). +Proof. +Admitted. + +(* ********** *) + +(* Exercise 08 *) + +Property bar : + forall (P : nat -> Prop) + (i : nat), + P i -> + forall j : nat, + P j. +Proof. +Abort. (* does not hold, see just below *) + +Theorem ex_falso_quodlibet_eg_False_revisited : + (forall (P : nat -> Prop) + (i : nat), + P i -> + forall j : nat, + P j) -> + 0 = 1. +Proof. +Abort. + +(* ********** *) + +(* Exercise 10 *) + +Property bar : (* stated again, but admitted this time... *) + forall (P : nat -> Prop) + (i : nat), + P i -> + forall j : nat, + P j. +Proof. +Admitted. (* ...to prove the following theorem as a corollary of bar *) + +Theorem ex_falso_quodlibet_eg_True_revisited : + forall m : nat, + exists n : nat, + m < n. +Proof. + Check (bar (fun m : nat => exists n : nat, m < n)). +Abort. + +(* ********** *) + +(* end of week-06_ex-falso-quodlibet.v *) diff --git a/cs3234/labs/week-06_miscellany.v b/cs3234/labs/week-06_miscellany.v new file mode 100644 index 0000000..8b54989 --- /dev/null +++ b/cs3234/labs/week-06_miscellany.v @@ -0,0 +1,320 @@ +(* week-06_miscellany.v *) +(* LPP 2024 - CS3234 2023-2024, Sem2 *) +(* Olivier Danvy *) +(* Version of 22 Feb 2024 *) + +(* ********** *) + +Require Import Arith Bool. + +(* ********** *) + +Lemma truism : + forall P : nat -> Prop, + (exists n : nat, + P n) -> + exists n : nat, + P n. +Proof. + intros P H_P. + exact H_P. + + Restart. + + intros P H_P. + destruct H_P as [n H_Pn]. + exists n. + exact H_Pn. + + Restart. + + intros P [n H_Pn]. + exists n. + exact H_Pn. +Qed. + +(* ***** *) + +Lemma other_truism : + forall P Q : nat -> Prop, + (exists n : nat, + P n /\ Q n) -> + exists m : nat, + P m \/ Q m. +Proof. + intros P Q [n [H_Pn H_Qn]]. + exists n. + left. + exact H_Pn. + + Restart. + + intros P Q [n [H_Pn H_Qn]]. + exists n. + right. + exact H_Qn. +Qed. + +(* ********** *) + +Lemma about_the_existential_quantifier_and_disjunction : + forall P Q : nat -> Prop, + (exists n : nat, P n \/ Q n) + <-> + ((exists n : nat, P n) + \/ + (exists n : nat, Q n)). +Proof. + intros P Q. + split. + - intros [n [H_P | H_Q]]. + + left. + exists n. + exact H_P. + + right. + exists n. + exact H_Q. + - intros [[n H_P] | [n H_Q]]. + + exists n. + left. + exact H_P. + + exists n. + right. + exact H_Q. +Qed. + +(* ********** *) + +Lemma about_the_universal_quantifier_and_conjunction : + forall P Q : nat -> Prop, + (forall n : nat, P n /\ Q n) + <-> + ((forall n : nat, P n) + /\ + (forall n : nat, Q n)). +Proof. + intros P Q. + split. + - intro H_PQ. + split. + + intro n. + destr + + Restart. + intros P Q. + split. + - intro H_PQ. + split. + + intro n. + destruct (H_PQ n) as [H_Pn _]. + exact H_Pn. + + intro n. + destruct (H_PQ n) as [_ H_Qn]. + exact H_Qn. + - intros [H_P H_Q] n. + exact (conj (H_P n) (H_Q n)). +Qed. + +(* ********** *) + +Definition specification_of_addition (add : nat -> nat -> nat) := + (forall m : nat, + add O m = m) + /\ + (forall n' m : nat, + add (S n') m = S (add n' m)). + +Definition specification_of_addition' (add : nat -> nat -> nat) := + forall n' m : nat, + add O m = m + /\ + add (S n') m = S (add n' m). + +Lemma about_two_universal_quantifiers_and_conjunction : + forall (P : nat -> Prop) + (Q : nat -> nat -> Prop), + ((forall j : nat, P j) + /\ + (forall i j : nat, Q i j)) + <-> + (forall i j : nat, P j /\ Q i j). +Proof. + intros P Q. + split. + - intros [H_P H_Q] i j. + split. + + exact (H_P j). + + exact (H_Q i j). + - intro H_PQ. + split. + + intro j. + destruct (H_PQ 0 j) as [H_Pj _]. + exact H_Pj. + + intros i j. + destruct (H_PQ i j) as [_ H_Qij]. + exact H_Qij. +Qed. + +Proposition the_two_specifications_of_addition_are_equivalent : + forall add : nat -> nat -> nat, + specification_of_addition add <-> specification_of_addition' add. +Proof. + intro add. + unfold specification_of_addition, specification_of_addition'. + Check (about_two_universal_quantifiers_and_conjunction + (fun m : nat => add 0 m = m) + (fun n' m : nat => add (S n') m = S (add n' m))). + exact (about_two_universal_quantifiers_and_conjunction + (fun m : nat => add 0 m = m) + (fun n' m : nat => add (S n') m = S (add n' m))). +Qed. + +(* ********** *) + +Lemma even_or_odd_dropped : + forall n : nat, + (exists q : nat, + n = 2 * q) + \/ + (exists q : nat, + n = S (2 * q)). +Proof. +Admitted. + +Lemma even_or_odd_lifted : + forall n : nat, + exists q : nat, + n = 2 * q + \/ + n = S (2 * q). +Proof. +Admitted. + +Proposition the_two_specifications_of_even_or_odd_are_equivalent : + forall n : nat, + (exists q : nat, + n = 2 * q + \/ + n = S (2 * q)) + <-> + ((exists q : nat, + n = 2 * q) + \/ + (exists q : nat, + n = S (2 * q))). +Proof. + intro n. + Check (about_the_existential_quantifier_and_disjunction + (fun q : nat => n = 2 * q) + (fun q : nat => n = S (2 * q))). + exact (about_the_existential_quantifier_and_disjunction + (fun q : nat => n = 2 * q) + (fun q : nat => n = S (2 * q))). +Qed. + +(* ********** *) + +Proposition factoring_and_distributing_a_forall_in_a_conclusion : + forall (P : nat -> Prop) + (Q : Prop), + (Q -> forall n : nat, P n) + <-> + (forall n : nat, + Q -> P n). +Proof. + intros P Q. + split. + - intros H_QP n H_Q. + exact (H_QP H_Q n). + - intros H_QP H_Q n. + exact (H_QP n H_Q). +Qed. + +(* ********** *) + +Proposition interplay_between_quantifiers_and_implication : + forall (P : nat -> Prop) + (Q : Prop), + (exists n : nat, P n -> Q) -> + (forall n : nat, P n) -> Q. +Proof. + intros P Q [n H_PnQ] H_P. + Check (H_PnQ (H_P n)). + exact (H_PnQ (H_P n)). +Qed. + +(* ********** *) + +Proposition interplay_between_implication_and_quantifiers : + forall (P : nat -> Prop) + (Q : Prop), + ((exists n : nat, P n) -> Q) -> + forall n : nat, P n -> Q. +Proof. + intros P Q H_PQ n H_Pn. + apply H_PQ. + exists n. + exact H_Pn. +Qed. + +(* ********** *) + +Proposition strengthening_X_in_the_conclusion : + forall A B C D X Y : Prop, + (A -> X) -> (B -> Y) -> (X -> C) -> (Y -> D) -> (X -> Y) -> A -> Y. +Proof. + intros A B C D X Y H_AX H_BY H_XC H_YD H_XY. +Abort. + +Proposition weakening_X_in_the_conclusion : + forall A B C D X Y : Prop, + (A -> X) -> (B -> Y) -> (X -> C) -> (Y -> D) -> (X -> Y) -> C -> Y. +Proof. + intros A B C D X Y H_AX H_BY H_XC H_YD H_XY. +Abort. + +Proposition strengthening_Y_in_the_conclusion : + forall A B C D X Y : Prop, + (A -> X) -> (B -> Y) -> (X -> C) -> (Y -> D) -> (X -> Y) -> X -> B. +Proof. + intros A B C D X Y H_AX H_BY H_XC H_YD H_XY. +Abort. + +Proposition weakening_Y_in_the_conclusion : + forall A B C D X Y : Prop, + (A -> X) -> (B -> Y) -> (X -> C) -> (Y -> D) -> (X -> Y) -> X -> D. +Proof. + intros A B C D X Y H_AX H_BY H_XC H_YD H_XY. +Abort. + +Proposition strengthening_X_in_a_premise : + forall A B C D X Y : Prop, + (A -> X) -> (B -> Y) -> (X -> C) -> (Y -> D) -> (A -> Y) -> X -> Y. +Proof. + intros A B C D X Y H_AX H_BY H_XC H_YD. +Abort. + +Proposition weakening_X_in_a_premise : + forall A B C D X Y : Prop, + (A -> X) -> (B -> Y) -> (X -> C) -> (Y -> D) -> (C -> Y) -> X -> Y. +Proof. + intros A B C D X Y H_AX H_BY H_XC H_YD. +Abort. + +Proposition strengthening_Y_in_a_premise : + forall A B C D X Y : Prop, + (A -> X) -> (B -> Y) -> (X -> C) -> (Y -> D) -> (X -> B) -> X -> Y. +Proof. + intros A B C D X Y H_AX H_BY H_XC H_YD. +Abort. + +Proposition weakening_Y_in_a_premise : + forall A B C D X Y : Prop, + (A -> X) -> (B -> Y) -> (X -> C) -> (Y -> D) -> (X -> D) -> X -> Y. +Proof. + intros A B C D X Y H_AX H_BY H_XC H_YD. +Abort. + +(* ********** *) + +(* end of week-06_miscellany.v *) diff --git a/cs3234/labs/week-06_soundness-and-completeness-of-equality-predicates.v b/cs3234/labs/week-06_soundness-and-completeness-of-equality-predicates.v new file mode 100644 index 0000000..dcd662c --- /dev/null +++ b/cs3234/labs/week-06_soundness-and-completeness-of-equality-predicates.v @@ -0,0 +1,784 @@ +(* week-06_soundness-and-completeness-of-equality-predicates.v *) +(* LPP 2024 - CS3234 2023-2024, Sem2 *) +(* Olivier Danvy *) +(* Version of 22 Feb 2024 *) + +(* ********** *) + +(* Paraphernalia: *) + +Ltac fold_unfold_tactic name := intros; unfold name; fold name; reflexivity. + +Require Import Arith Bool. + +(* ********** *) + +Check Bool.eqb. (* : bool -> bool -> bool *) + +Check eqb. (* : bool -> bool -> bool *) + +Search (eqb _ _ = true -> _ = _). +(* eqb_prop: forall a b : bool, eqb a b = true -> a = b *) + +Search (eqb _ _ = true). +(* eqb_reflx: forall b : bool, eqb b b = true *) + +Theorem soundness_of_equality_over_booleans : + forall b1 b2 : bool, + eqb b1 b2 = true -> b1 = b2. +Proof. + exact eqb_prop. + + Restart. + + intros [ | ] [ | ]. + - intros _. + reflexivity. + - unfold eqb. + intro H_absurd. + discriminate H_absurd. + - unfold eqb. + intro H_absurd. + exact H_absurd. + - intros _. + reflexivity. +Qed. + +Theorem completeness_of_equality_over_booleans : + forall b1 b2 : bool, + b1 = b2 -> eqb b1 b2 = true. +Proof. + intros b1 b2 H_b1_b2. + rewrite <- H_b1_b2. + Search (eqb _ _ = true). + Check (eqb_reflx b1). + exact (eqb_reflx b1). + + Restart. + + intros [ | ] [ | ]. + - intros _. + unfold eqb. + reflexivity. + - intros H_absurd. + discriminate H_absurd. + - intros H_absurd. + discriminate H_absurd. + - intros _. + unfold eqb. + reflexivity. +Qed. + +Corollary soundness_of_equality_over_booleans_the_remaining_case : + forall b1 b2 : bool, + eqb b1 b2 = false -> b1 <> b2. +Proof. + intros b1 b2 H_eqb_b1_b2. + unfold not. + intros H_eq_b1_b2. + Check (completeness_of_equality_over_booleans b1 b2 H_eq_b1_b2). + rewrite -> (completeness_of_equality_over_booleans b1 b2 H_eq_b1_b2) in H_eqb_b1_b2. + discriminate H_eqb_b1_b2. +Qed. + +Corollary completeness_of_equality_over_booleans_the_remaining_case : + forall b1 b2 : bool, + b1 <> b2 -> eqb b1 b2 = false. +Proof. + intros b1 b2 H_neq_b1_b2. + unfold not in H_neq_b1_b2. + Search (not (_ = true) -> _ = false). + Check (not_true_is_false (eqb b1 b2)). + apply (not_true_is_false (eqb b1 b2)). + unfold not. + intro H_eqb_b1_b2. + Check (soundness_of_equality_over_booleans b1 b2 H_eqb_b1_b2). + Check (H_neq_b1_b2 (soundness_of_equality_over_booleans b1 b2 H_eqb_b1_b2)). + contradiction (H_neq_b1_b2 (soundness_of_equality_over_booleans b1 b2 H_eqb_b1_b2)). +(* Or alternatively: + exact (H_neq_b1_b2 (soundness_of_equality_over_booleans b1 b2 H_eqb_b1_b2)). +*) +Qed. + +Check Bool.eqb_eq. +(* eqb_eq : forall x y : bool, Is_true (eqb x y) -> x = y *) + +Search (eqb _ _ = true). +(* eqb_true_iff: forall a b : bool, eqb a b = true <-> a = b *) + +Theorem soundness_and_completeness_of_equality_over_booleans : + forall b1 b2 : bool, + eqb b1 b2 = true <-> b1 = b2. +Proof. + exact eqb_true_iff. + + Restart. + + intros b1 b2. + split. + - exact (soundness_of_equality_over_booleans b1 b2). + - exact (completeness_of_equality_over_booleans b1 b2). +Qed. + +(* ***** *) + +(* user-defined: *) + +Definition eqb_bool (b1 b2 : bool) : bool := + match b1 with + true => + match b2 with + true => + true + | false => + false + end + | false => + match b2 with + true => + false + | false => + true + end + end. + +Theorem soundness_of_eqb_bool : + forall b1 b2 : bool, + eqb_bool b1 b2 = true -> + b1 = b2. +Proof. + intros [ | ] [ | ]. + - intros _. + reflexivity. + - unfold eqb_bool. + intros H_absurd. + discriminate H_absurd. + - unfold eqb_bool. + intros H_absurd. + exact H_absurd. + - intros _. + reflexivity. +Qed. + +Theorem completeness_of_eqb_bool : + forall b1 b2 : bool, + b1 = b2 -> + eqb_bool b1 b2 = true. +Proof. + intros [ | ] [ | ]. + - intros _. + reflexivity. + - intros H_absurd. + discriminate H_absurd. + - intros H_absurd. + unfold eqb_bool. + exact H_absurd. + - intros _. + unfold eqb_bool. + reflexivity. +Qed. + +(* ********** *) + +Check Nat.eqb. (* : nat -> nat -> bool *) + +Check beq_nat. (* : nat -> nat -> bool *) + +Search (beq_nat _ _ = true -> _ = _). +(* beq_nat_true: forall n m : nat, (n =? m) = true -> n = m *) + +Search (beq_nat _ _ = true). + +(* Nat.eqb_eq: forall n m : nat, (n =? m) = true <-> n = m *) + +Theorem soundness_and_completeness_of_equality_over_natural_numbers : + forall n1 n2 : nat, + n1 =? n2 = true <-> n1 = n2. +Proof. + exact Nat.eqb_eq. +Qed. + +(* ***** *) + +(* user-defined: *) + +Fixpoint eqb_nat (n1 n2 : nat) : bool := + match n1 with + O => + match n2 with + O => + true + | S n2' => + false + end + | S n1' => + match n2 with + O => + false + | S n2' => + eqb_nat n1' n2' + end + end. + +Lemma fold_unfold_eqb_nat_O : + forall n2 : nat, + eqb_nat O n2 = + match n2 with + O => + true + | S n2' => + false + end. +Proof. + fold_unfold_tactic eqb_nat. +Qed. + +Lemma fold_unfold_eqb_nat_S : + forall n1' n2 : nat, + eqb_nat (S n1') n2 = + match n2 with + O => + false + | S n2' => + eqb_nat n1' n2' + end. +Proof. + fold_unfold_tactic eqb_nat. +Qed. + +Theorem soundness_of_eqb_nat : + forall n1 n2 : nat, + eqb_nat n1 n2 = true -> + n1 = n2. +Proof. + intro n1. + induction n1 as [ | n1' IHn1']. + - intros [ | n2']. + + intros _. + reflexivity. + + rewrite -> fold_unfold_eqb_nat_O. + intro H_absurd. + discriminate H_absurd. + - intros [ | n2']. + + rewrite -> fold_unfold_eqb_nat_S. + intro H_absurd. + discriminate H_absurd. + + rewrite -> fold_unfold_eqb_nat_S. + intro H_n1'_n2'. + Check (IHn1' n2' H_n1'_n2'). + rewrite -> (IHn1' n2' H_n1'_n2'). + reflexivity. +Qed. + +Theorem completeness_of_eqb_nat : + forall n1 n2 : nat, + n1 = n2 -> + eqb_nat n1 n2 = true. +Proof. + intro n1. + induction n1 as [ | n1' IHn1']. + - intros [ | n2']. + + intros _. + rewrite -> fold_unfold_eqb_nat_O. + reflexivity. + + intro H_absurd. + discriminate H_absurd. + - intros [ | n2']. + + intro H_absurd. + discriminate H_absurd. + + rewrite -> fold_unfold_eqb_nat_S. + intro H_Sn1'_Sn2'. + injection H_Sn1'_Sn2' as H_n1'_n2'. + Check (IHn1' n2' H_n1'_n2'). + rewrite -> (IHn1' n2' H_n1'_n2'). + reflexivity. +Qed. + +(* ********** *) + +Lemma from_one_equivalence_to_two_implications : + forall (V : Type) + (eqb_V : V -> V -> bool), + (forall v1 v2 : V, + eqb_V v1 v2 = true <-> v1 = v2) -> + (forall v1 v2 : V, + eqb_V v1 v2 = true -> v1 = v2) + /\ + (forall v1 v2 : V, + v1 = v2 -> eqb_V v1 v2 = true). +Proof. + intros V eqb_V H_eqv. + split. + - intros v1 v2 H_eqb. + destruct (H_eqv v1 v2) as [H_key _]. + exact (H_key H_eqb). + - intros v1 v2 H_eq. + destruct (H_eqv v1 v2) as [_ H_key]. + exact (H_key H_eq). +Qed. + +(* ********** *) + +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. + +Theorem soundness_of_equality_over_optional_values : + forall (V : Type) + (eqb_V : V -> V -> bool), + (forall v1 v2 : V, + eqb_V v1 v2 = true -> v1 = v2) -> + forall ov1 ov2 : option V, + eqb_option V eqb_V ov1 ov2 = true -> + ov1 = ov2. +Proof. + intros V eqb_V S_eqb_V [v1 | ] [v2 | ] H_eqb. + - unfold eqb_option in H_eqb. + Check (S_eqb_V v1 v2 H_eqb). + rewrite -> (S_eqb_V v1 v2 H_eqb). + reflexivity. + - unfold eqb_option in H_eqb. + discriminate H_eqb. + - unfold eqb_option in H_eqb. + discriminate H_eqb. + - reflexivity. +Qed. + +Theorem completeness_of_equality_over_optional_values : + forall (V : Type) + (eqb_V : V -> V -> bool), + (forall v1 v2 : V, + v1 = v2 -> eqb_V v1 v2 = true) -> + forall ov1 ov2 : option V, + ov1 = ov2 -> + eqb_option V eqb_V ov1 ov2 = true. +Proof. + intros V eqb_V C_eqb_V ov1 ov2 H_eq. + rewrite -> H_eq. + case ov1 as [v1 | ]. + - case ov2 as [v2 | ]. + -- unfold eqb_option. + Check (eq_refl v2). + Check (C_eqb_V v2 v2 (eq_refl v2)). + exact (C_eqb_V v2 v2 (eq_refl v2)). + -- discriminate H_eq. + - case ov2 as [v2 | ]. + -- discriminate H_eq. + -- unfold eqb_option. + reflexivity. +Qed. + +Theorem soundness_and_completeness_of_equality_over_optional_values : + forall (V : Type) + (eqb_V : V -> V -> bool), + (forall v1 v2 : V, + eqb_V v1 v2 = true <-> v1 = v2) -> + forall ov1 ov2 : option V, + eqb_option V eqb_V ov1 ov2 = true <-> ov1 = ov2. +Proof. + intros V eqb_V SC_eqb_V. + Check (from_one_equivalence_to_two_implications V eqb_V SC_eqb_V). + destruct (from_one_equivalence_to_two_implications V eqb_V SC_eqb_V) as [S_eqb_V C_eqb_V]. + intros ov1 ov2. + split. + - exact (soundness_of_equality_over_optional_values V eqb_V S_eqb_V ov1 ov2). + - exact (completeness_of_equality_over_optional_values V eqb_V C_eqb_V ov1 ov2). +Qed. + +(* ********** *) + +Definition eqb_pair (V : Type) (eqb_V : V -> V -> bool) (W : Type) (eqb_W : W -> W -> bool) (p1 p2 : V * W) : bool := + let (v1, w1) := p1 in + let (v2, w2) := p2 in + eqb_V v1 v2 && eqb_W w1 w2. + +Theorem soundness_of_equality_over_pairs : + forall (V : Type) + (eqb_V : V -> V -> bool), + (forall v1 v2 : V, + eqb_V v1 v2 = true -> v1 = v2) -> + forall (W : Type) + (eqb_W : W -> W -> bool), + (forall w1 w2 : W, + eqb_W w1 w2 = true -> w1 = w2) -> + forall p1 p2 : V * W, + eqb_pair V eqb_V W eqb_W p1 p2 = true -> + p1 = p2. +Proof. + intros V eqb_V S_eqb_V W eqb_W S_eqb_W [v1 w1] [v2 w2] H_eqb. + unfold eqb_pair in H_eqb. + Search (_ && _ = true -> _ /\ _). + Check (andb_prop (eqb_V v1 v2) (eqb_W w1 w2)). + Check (andb_prop (eqb_V v1 v2) (eqb_W w1 w2) H_eqb). + destruct (andb_prop (eqb_V v1 v2) (eqb_W w1 w2) H_eqb) as [H_eqb_V H_eqb_W]. + Check (S_eqb_V v1 v2 H_eqb_V). + rewrite -> (S_eqb_V v1 v2 H_eqb_V). + rewrite -> (S_eqb_W w1 w2 H_eqb_W). + reflexivity. +Qed. + +Theorem completeness_of_equality_over_pairs : + forall (V : Type) + (eqb_V : V -> V -> bool), + (forall v1 v2 : V, + v1 = v2 -> eqb_V v1 v2 = true) -> + forall (W : Type) + (eqb_W : W -> W -> bool), + (forall w1 w2 : W, + w1 = w2 -> eqb_W w1 w2 = true) -> + forall p1 p2 : V * W, + p1 = p2 -> + eqb_pair V eqb_V W eqb_W p1 p2 = true. +Proof. + intros V eqb_V S_eqb_V W eqb_W S_eqb_W [v1 w1] [v2 w2] H_eq. + unfold eqb_pair. + injection H_eq as H_eq_V H_eq_W. + Check (S_eqb_V v1 v2 H_eq_V). + rewrite -> (S_eqb_V v1 v2 H_eq_V). + rewrite -> (S_eqb_W w1 w2 H_eq_W). + unfold andb. + reflexivity. +Qed. + +Theorem soundness_and_completeness_of_equality_over_pairs : + forall (V : Type) + (eqb_V : V -> V -> bool), + (forall v1 v2 : V, + eqb_V v1 v2 = true <-> v1 = v2) -> + forall (W : Type) + (eqb_W : W -> W -> bool), + (forall w1 w2 : W, + eqb_W w1 w2 = true <-> w1 = w2) -> + forall p1 p2 : V * W, + eqb_pair V eqb_V W eqb_W p1 p2 = true <-> p1 = p2. +Proof. + intros V eqb_V SC_eqb_V. + Check (from_one_equivalence_to_two_implications V eqb_V SC_eqb_V). + destruct (from_one_equivalence_to_two_implications V eqb_V SC_eqb_V) as [S_eqb_V C_eqb_V]. + intros W eqb_W SC_eqb_W. + Check (from_one_equivalence_to_two_implications W eqb_W SC_eqb_W). + destruct (from_one_equivalence_to_two_implications W eqb_W SC_eqb_W) as [S_eqb_W C_eqb_W]. + intros p1 p2. + split. + - exact (soundness_of_equality_over_pairs V eqb_V S_eqb_V W eqb_W S_eqb_W p1 p2). + - exact (completeness_of_equality_over_pairs V eqb_V C_eqb_V W eqb_W C_eqb_W p1 p2). +Qed. + +(* ********** *) + +Inductive binary_tree (V : Type) : Type := + Leaf : V -> binary_tree V +| Node : binary_tree V -> binary_tree V -> binary_tree V. + +Fixpoint eqb_binary_tree (V : Type) (eqb_V : V -> V -> bool) (t1 t2 : binary_tree V) : bool := + match t1 with + Leaf _ v1 => + match t2 with + Leaf _ v2 => + eqb_V v1 v2 + | Node _ t11 t12 => + false + end + | Node _ t11 t12 => + match t2 with + Leaf _ v2 => + false + | Node _ t21 t22 => + eqb_binary_tree V eqb_V t11 t21 + && + eqb_binary_tree V eqb_V t12 t22 + end + end. + +Lemma fold_unfold_eqb_binary_tree_Leaf : + forall (V : Type) + (eqb_V : V -> V -> bool) + (v1 : V) + (t2 : binary_tree V), + eqb_binary_tree V eqb_V (Leaf V v1) t2 = + match t2 with + Leaf _ v2 => + eqb_V v1 v2 + | Node _ t11 t12 => + false + end. +Proof. + fold_unfold_tactic eqb_binary_tree. +Qed. + +Lemma fold_unfold_eqb_binary_tree_Node : + forall (V : Type) + (eqb_V : V -> V -> bool) + (t11 t12 t2 : binary_tree V), + eqb_binary_tree V eqb_V (Node V t11 t12) t2 = + match t2 with + Leaf _ v2 => + false + | Node _ t21 t22 => + eqb_binary_tree V eqb_V t11 t21 + && + eqb_binary_tree V eqb_V t12 t22 + end. +Proof. + fold_unfold_tactic eqb_binary_tree. +Qed. + +Theorem soundness_of_equality_over_binary_trees : + forall (V : Type) + (eqb_V : V -> V -> bool), + (forall v1 v2 : V, + eqb_V v1 v2 = true -> v1 = v2) -> + forall t1 t2 : binary_tree V, + eqb_binary_tree V eqb_V t1 t2 = true -> + t1 = t2. +Proof. + intros V eqb_V S_eqb_V t1. + induction t1 as [v1 | t11 IHt11 t12 IHt12]. + - intros [v2 | t21 t22] H_eqb. + -- rewrite -> (fold_unfold_eqb_binary_tree_Leaf V eqb_V v1 (Leaf V v2)) in H_eqb. + Check (S_eqb_V v1 v2 H_eqb). + rewrite -> (S_eqb_V v1 v2 H_eqb). + reflexivity. + -- rewrite -> (fold_unfold_eqb_binary_tree_Leaf V eqb_V v1 (Node V t21 t22)) in H_eqb. + discriminate H_eqb. + - intros [v2 | t21 t22] H_eqb. + -- rewrite -> (fold_unfold_eqb_binary_tree_Node V eqb_V t11 t12 (Leaf V v2)) in H_eqb. + discriminate H_eqb. + -- rewrite -> (fold_unfold_eqb_binary_tree_Node V eqb_V t11 t12 (Node V t21 t22)) in H_eqb. + Search (_ && _ = true -> _ /\ _). + Check (andb_prop (eqb_binary_tree V eqb_V t11 t21) (eqb_binary_tree V eqb_V t12 t22)). + Check (andb_prop (eqb_binary_tree V eqb_V t11 t21) (eqb_binary_tree V eqb_V t12 t22) H_eqb). + destruct (andb_prop (eqb_binary_tree V eqb_V t11 t21) (eqb_binary_tree V eqb_V t12 t22) H_eqb) as [H_eqb_1 H_eqb_2]. + Check (IHt11 t21 H_eqb_1). + rewrite -> (IHt11 t21 H_eqb_1). + rewrite -> (IHt12 t22 H_eqb_2). + reflexivity. +Qed. + +Theorem completeness_of_equality_over_binary_trees : + forall (V : Type) + (eqb_V : V -> V -> bool), + (forall v1 v2 : V, + v1 = v2 -> eqb_V v1 v2 = true) -> + forall t1 t2 : binary_tree V, + t1 = t2 -> + eqb_binary_tree V eqb_V t1 t2 = true. +Proof. + intros V eqb_V C_eqb_V t1. + induction t1 as [v1 | t11 IHt11 t12 IHt12]. + - intros [v2 | t21 t22] H_eq. + -- rewrite -> (fold_unfold_eqb_binary_tree_Leaf V eqb_V v1 (Leaf V v2)). + injection H_eq as H_eq_V. + Check (C_eqb_V v1 v2). + Check (C_eqb_V v1 v2 H_eq_V). + exact (C_eqb_V v1 v2 H_eq_V). + -- discriminate H_eq. + - intros [v2 | t21 t22] H_eq. + -- discriminate H_eq. + -- rewrite -> (fold_unfold_eqb_binary_tree_Node V eqb_V t11 t12 (Node V t21 t22)). + injection H_eq as H_eq_1 H_eq_2. + Check (IHt11 t21 H_eq_1). + rewrite -> (IHt11 t21 H_eq_1). + Search (true && _ = _). + rewrite -> (andb_true_l (eqb_binary_tree V eqb_V t12 t22)). + exact (IHt12 t22 H_eq_2). +Qed. + +Theorem soundness_and_completeness_of_equality_over_binary_trees : + forall (V : Type) + (eqb_V : V -> V -> bool), + (forall v1 v2 : V, + eqb_V v1 v2 = true <-> v1 = v2) -> + forall t1 t2 : binary_tree V, + eqb_binary_tree V eqb_V t1 t2 = true <-> t1 = t2. +Proof. + intros V eqb_V SC_eqb_V t1 t2. + Check (from_one_equivalence_to_two_implications V eqb_V SC_eqb_V). + destruct (from_one_equivalence_to_two_implications V eqb_V SC_eqb_V) as [S_eqb_V C_eqb_V]. + split. + - exact (soundness_of_equality_over_binary_trees V eqb_V S_eqb_V t1 t2). + - exact (completeness_of_equality_over_binary_trees V eqb_V C_eqb_V t1 t2). + + Restart. + + intros V eqb_V SC_eqb_V t1. + induction t1 as [v1 | t11 IHt11 t12 IHt12]. + - intros [v2 | t21 t22]. + + rewrite -> (fold_unfold_eqb_binary_tree_Leaf V eqb_V v1 (Leaf V v2)). + split. + * intro H_eqb_V. + destruct (from_one_equivalence_to_two_implications V eqb_V SC_eqb_V) as [S_eqb_V _]. + rewrite -> (S_eqb_V v1 v2 H_eqb_V). + reflexivity. + * intro H_eq. + injection H_eq as H_eq. + destruct (from_one_equivalence_to_two_implications V eqb_V SC_eqb_V) as [_ C_eqb_V]. + exact (C_eqb_V v1 v2 H_eq). + + rewrite -> (fold_unfold_eqb_binary_tree_Leaf V eqb_V v1 (Node V t21 t22)). + split. + * intro H_absurd. + discriminate H_absurd. + * intro H_absurd. + discriminate H_absurd. + - intros [v2 | t21 t22]. + + rewrite -> (fold_unfold_eqb_binary_tree_Node V eqb_V t11 t12 (Leaf V v2)). + split. + * intro H_absurd. + discriminate H_absurd. + * intro H_absurd. + discriminate H_absurd. + + rewrite -> (fold_unfold_eqb_binary_tree_Node V eqb_V t11 t12 (Node V t21 t22)). + split. + * intro H_eqb. + destruct (andb_prop (eqb_binary_tree V eqb_V t11 t21) (eqb_binary_tree V eqb_V t12 t22) H_eqb) as [H_eqb_1 H_eqb_2]. + destruct (IHt11 t21) as [H_key1 _]. + destruct (IHt12 t22) as [H_key2 _]. + rewrite -> (H_key1 H_eqb_1). + rewrite -> (H_key2 H_eqb_2). + reflexivity. + * intro H_eq. + injection H_eq as H_eq_1 H_eq_2. + destruct (IHt11 t21) as [_ H_key1]. + destruct (IHt12 t22) as [_ H_key2]. + rewrite -> (H_key1 H_eq_1). + rewrite -> (andb_true_l (eqb_binary_tree V eqb_V t12 t22)). + exact (H_key2 H_eq_2). +Qed. + +(* ********** *) + +(* pilfering from the String library: *) + +Require Import String Ascii. + +Print string. + +Check "foo"%string. + +Definition eqb_char (c1 c2 : ascii) : bool := + match c1 with + Ascii b11 b12 b13 b14 b15 b16 b17 b18 => + match c2 with + Ascii b21 b22 b23 b24 b25 b26 b27 b28 => + eqb_bool b11 b21 && eqb_bool b12 b22 && eqb_bool b13 b23 && eqb_bool b14 b24 && eqb_bool b15 b25 && eqb_bool b16 b26 && eqb_bool b17 b27 && eqb_bool b18 b28 + end + end. + +Proposition soundness_of_eqb_char : + forall c1 c2 : ascii, + eqb_char c1 c2 = true -> c1 = c2. +Proof. +Admitted. + +Proposition completeness_of_eqb_char : + forall c1 c2 : ascii, + c1 = c2 -> eqb_char c1 c2 = true. +Proof. +Admitted. + +Fixpoint eqb_string (c1s c2s : string) : bool := + match c1s with + String.EmptyString => + match c2s with + String.EmptyString => + true + | String.String c2 c2s' => + false + end + | String.String c1 c1s' => + match c2s with + String.EmptyString => + true + | String.String c2 c2s' => + eqb_char c1 c2 && eqb_string c1s' c2s' + end + end. + +Lemma fold_unfold_eqb_string_Empty : + forall c2s : string, + eqb_string String.EmptyString c2s = + match c2s with + String.EmptyString => + true + | String.String c2 c2s' => + false + end. +Proof. + fold_unfold_tactic eqb_string. +Qed. + +Lemma fold_unfold_eqb_string_String : + forall (c1 : ascii) + (c1s' c2s : string), + eqb_string (String.String c1 c1s') c2s = + match c2s with + String.EmptyString => + true + | String.String c2 c2s' => + eqb_char c1 c2 && eqb_string c1s' c2s' + end. +Proof. + fold_unfold_tactic eqb_string. +Qed. + +Proposition soundness_of_eqb_string : + forall c1s c2s : string, + eqb_string c1s c2s = true -> c1s = c2s. +Proof. +Admitted. + +Proposition completeness_of_eqb_string : + forall c1s c2s : string, + c1s = c2s -> eqb_string c1s c2s = true. +Proof. +Admitted. + +(* ********** *) + +Inductive funky_tree : Type := + Nat : nat -> funky_tree +| Bool : bool -> funky_tree +| String : string -> funky_tree +| Singleton : funky_tree -> funky_tree +| Pair : funky_tree -> funky_tree -> funky_tree +| Triple : funky_tree -> funky_tree -> funky_tree -> funky_tree. + +(* ***** *) + +(* A silly proposition, just to get a feel about how to destructure a value of type funky_tree: *) + +Proposition identity_over_funky_tree : + forall e : funky_tree, + e = e. +Proof. + intro e. + case e as [n | b | s | e1 | e1 e2 | e1 e2 e3] eqn:H_e. + - reflexivity. + - reflexivity. + - reflexivity. + - reflexivity. + - reflexivity. + - reflexivity. +Qed. + +(* ***** *) + +(* Exercise: implement eqb_funky_tree and prove its soundness and completeness. *) + +(* ********** *) + +(* end of week-06_soundness-and-completeness-of-equality-predicates.v *) diff --git a/cs3234/labs/week-07_isometries3.v b/cs3234/labs/week-07_isometries3.v new file mode 100644 index 0000000..7ec4246 --- /dev/null +++ b/cs3234/labs/week-07_isometries3.v @@ -0,0 +1,514 @@ +(* week-07_isometries3.v *) +(* LPP 2024 - CS3234 2023-2024, Sem2 *) +(* Olivier Danvy *) +(* Version of 08 Mar 2024 *) + +(* ********** *) + +(* A formal study of isometries of the equilateral triangle, *) +(* after Chantal Keller, Damien Pous and Sylvain Chevillard. *) + +(* ********** *) + +Inductive Rotation : Type := + R000 : Rotation (* 0 degrees (identity) *) +| R120 : Rotation (* 120 degrees *) +| R240 : Rotation. (* 240 degrees *) + +(* ********** *) + +(* Performing two rotations in a row, clockwise. *) +(* "RaR" stands for "a Rotation after a Rotation" *) + +Definition RaR (r2 r1: Rotation) : Rotation := + match r1 with + R000 => match r2 with + R000 => R000 + | R120 => R120 + | R240 => R240 + end + | R120 => match r2 with + R000 => R120 + | R120 => R240 + | R240 => R000 + end + | R240 => match r2 with + R000 => R240 + | R120 => R000 + | R240 => R120 + end + end. + +(* Some properties: *) + +Proposition R000_is_neutral_for_RaR_on_the_left : + forall r : Rotation, + RaR R000 r = r. +Proof. + intros [ | | ]. + - unfold RaR. + reflexivity. + - unfold RaR. + reflexivity. + - unfold RaR. + reflexivity. +Qed. + +Proposition R000_is_neutral_for_RaR_on_the_right : + forall r : Rotation, + RaR r R000 = r. +Proof. + intros [ | | ]; unfold RaR. + reflexivity. + reflexivity. + reflexivity. +Qed. + +Proposition RaR_is_commutative : + forall r1 r2 : Rotation, + RaR r2 r1 = RaR r1 r2. +Proof. + intros [ | | ] [ | | ]; unfold RaR; reflexivity. + Restart. + intros [ | | ] [ | | ]; reflexivity. +Qed. + +Proposition RaR_is_associative : + forall r1 r2 r3 : Rotation, + RaR (RaR r3 r2) r1 = RaR r3 (RaR r2 r1). +Proof. + intros [ | | ] [ | | ] [ | | ]; reflexivity. +Abort. + +Proposition RaR_is_nilpotent_with_order_3 : + forall r : Rotation, + RaR (RaR r r) (RaR r r) = r. +Proof. + intros [ | | ]; unfold RaR; reflexivity. +Qed. + +(* Abort. *) + +(* ********** *) + +(* The following symmetries are indexed by the invariant vertex: *) + +Inductive Reflection : Type := + S_NN : Reflection (* North, at the top *) +| S_SW : Reflection (* South-West, at the bottom left *) +| S_SE : Reflection. (* South-East, at the bottom right *) + +(* These reflections are symmetries here. *) + +(* Performing two reflections in a row. *) +(* "SaS" stands for "a Symmetry after a Symmetry" *) + +Definition SaS (s2 s1 : Reflection) : Rotation := + match s1 with + S_NN => match s2 with + S_NN => R000 + | S_SW => R120 + | S_SE => R240 + end + | S_SW => match s2 with + S_NN => R240 + | S_SW => R000 + | S_SE => R120 + end + | S_SE => match s2 with + S_NN => R120 + | S_SW => R240 + | S_SE => R000 + end + end. + +(* is SaS commutative? *) +Proposition SaS_is_not_commutative : + exists s1 s2 : Reflection, + SaS s1 s2 <> SaS s2 s1. +Proof. + exists S_NN, S_SW. + unfold SaS. + unfold not. + intro H_absurd. + discriminate H_absurd. +Qed. +(* is SaS associative? *) + +(* is SaS nilpotent? *) + +(* ********** *) + +(* Performing a rotation and then a reflection in a row. *) +(* "SaR" stands for "a Symmetry after a Rotation" *) + +Definition SaR (s : Reflection) (r : Rotation) : Reflection := + match r with + R000 => match s with + S_NN => S_NN + | S_SW => S_SW + | S_SE => S_SE + end + | R120 => match s with + S_NN => S_SE + | S_SW => S_NN + | S_SE => S_SW + end + | R240 => match s with + S_NN => S_SW + | S_SW => S_SE + | S_SE => S_NN + end + end. + +(* ********** *) + +(* Performing a reflection and then a rotation in a row. *) +(* "RaS" stands for "a Rotation after a Symmetry" *) + +Definition RaS (r : Rotation) (s : Reflection) : Reflection := + match s with + S_NN => match r with + R000 => S_NN + | R120 => S_SW + | R240 => S_SE + end + | S_SW => match r with + R000 => S_SW + | R120 => S_SE + | R240 => S_NN + end + | S_SE => match r with + R000 => S_SE + | R120 => S_NN + | R240 => S_SW + end + end. + +(* ********** *) + +Inductive Isomorphism : Type := +| IR : Rotation -> Isomorphism +| IS : Reflection -> Isomorphism. + +(* Identity: *) + +Definition Id : Isomorphism := IR R000. + +(* Composition: *) + +Definition C (i2 i1 : Isomorphism) : Isomorphism := + match i1 with + IR r1 => match i2 with + IR r2 => IR (RaR r2 r1) + | IS s2 => IS (SaR s2 r1) + end + | IS s1 => match i2 with + IR r2 => IS (RaS r2 s1) + | IS s2 => IR (SaS s2 s1) + end + end. + +Proposition Id_is_neutral_on_the_left_of_C : + forall i : Isomorphism, + C Id i = i. +Proof. + intros [[ | | ] | [ | | ]]; reflexivity. +Qed. + +Proposition Id_is_neutral_on_the_right_of_C : + forall i : Isomorphism, + C i Id = i. +Proof. + intros [[ | | ] | [ | | ]]; reflexivity. +Qed. + +Proposition C_is_associative : + forall i1 i2 i3 : Isomorphism, + C i3 (C i2 i1) = C (C i3 i2) i1. +Proof. + intros [[ | | ] | [ | | ]] + [[ | | ] | [ | | ]] + [[ | | ] | [ | | ]] ; reflexivity. +Abort. + +Lemma composing_an_isomorphism_is_injective_on_the_right : + forall i x y : Isomorphism, + C i x = C i y -> x = y. +Proof. + intros [[ | | ] | [ | | ]] + [[ | | ] | [ | | ]] + [[ | | ] | [ | | ]]; unfold C, RaR, SaR, RaS, SaS. + - intro H. + reflexivity. + - intro H_absurd. + discriminate H_absurd. + Restart. + intros [[ | | ] | [ | | ]] + [[ | | ] | [ | | ]] + [[ | | ] | [ | | ]]; + intro H; (reflexivity || discriminate H). + + + +Qed. + +Lemma composing_an_isomorphism_is_injective_on_the_left : + forall i x y : Isomorphism, + C x i = C y i -> x = y. +Proof. + intros [[ | | ] | [ | | ]] + [[ | | ] | [ | | ]] + [[ | | ] | [ | | ]]; + intro H; (reflexivity || discriminate H). +Qed. + +Lemma composing_an_isomorphism_is_surjective_on_the_right : + forall i2 i3 : Isomorphism, + exists i1 : Isomorphism, + C i2 i1 = i3. +Proof. + intros [[ | | ] | [ | | ]] + [[ | | ] | [ | | ]]. + - exists Id. + reflexivity. + - exists (IR R120). + reflexivity. + - exists (IR R240). + reflexivity. + Restart. + intros [[ | | ] | [ | | ]] + [[ | | ] | [ | | ]]; + ( + (exists (IR R000); reflexivity ) + || + (exists (IR R120); reflexivity ) + || + (exists (IR R240); reflexivity ) + || + (exists (IS S_NN); reflexivity ) + || + (exists (IS S_SW); reflexivity ) + || + (exists (IS S_SE); reflexivity ) + ). + Show Proof. +Qed. + +Lemma composing_an_isomorphism_is_surjective_on_the_left : + forall i1 i3 : Isomorphism, + exists i2 : Isomorphism, + C i2 i1 = i3. +Proof. + intros [[ | | ] | [ | | ]] + [[ | | ] | [ | | ]]; + ( + (exists (IR R000); reflexivity ) + || + (exists (IR R120); reflexivity ) + || + (exists (IR R240); reflexivity ) + || + (exists (IS S_NN); reflexivity ) + || + (exists (IS S_SW); reflexivity ) + || + (exists (IS S_SE); reflexivity ) + ). +Qed. + +Proposition C_over_rotations_is_nilpotent_with_order_3 : + forall r : Rotation, + C (C (IR r) (IR r)) (IR r) = Id. +Proof. +Abort. + +Proposition C_over_symmetries_is_nilpotent_with_order_2 : + forall s : Reflection, + C (IS s) (IS s) = Id. +Proof. + intros [ | | ]; unfold C, SaS, Id; reflexivity. +Qed. + +(* Proposition C_is_nilpotent_with_order_??? : *) +(* forall i : Isomorphism, *) +(* ... *) +(* Proof. *) +(* Abort. *) + +(* ********** *) + +(* Let us now introduce the vertices: *) + +Inductive Vertex : Type := (* enumerated clockwise *) + NN : Vertex +| SW : Vertex +| SE : Vertex. + +(* And let us define the effect of applying an isomorphism + to a vertex: *) + +Definition A (i : Isomorphism) (v : Vertex) : Vertex := + match i with + IR r => match r with + R000 => match v with + NN => NN + | SW => SW + | SE => SE + end + | R120 => match v with + NN => SW + | SW => SE + | SE => NN + end + | R240 => match v with + NN => SE + | SE => SW + | SW => NN + end + end + | IS s => match s with + S_NN => match v with + NN => NN + | SW => SE + | SE => SW + end + | S_SE => match v with + NN => SW + | SW => NN + | SE => SE + end + | S_SW => match v with + NN => SE + | SW => SW + | SE => NN + end + end + end. + +(* +Proposition A_is_equivalent_to_C : + forall (i1 i2 : Isomorphism) (v : Vertex), + A i2 (A i1 v) = A (C i2 i1) v. +Proof. +Abort. + +Proposition applying_an_isomorphism_is_injective : + forall (i : Isomorphism) (v1 v2 : Vertex), + (A i v1 = A i v2) -> v1 = v2. +Proof. +Abort. + +Proposition applying_an_isomorphism_is_surjective : + forall (i : Isomorphism) (v2 : Vertex), + exists v1 : Vertex, + A i v1 = v2. +Proof. +Abort. +*) + +(* ********** *) + +(* Intensional equality: + two isomorphisms are equal + iff + they are are constructed alike. + *) + +Definition intensional_equality (i1 i2: Isomorphism) : Prop := + i1 = i2. + +(* Extensional equality: + two isomorphisms are equal + iff + their graphs are the same. + *) + +Definition extensional_equality (i1 i2: Isomorphism) : Prop := + forall v : Vertex, + A i1 v = A i2 v. + +(* The two notions of equalities coincide: *) + +Proposition the_two_equalities_are_the_same : + forall i1 i2 : Isomorphism, + intensional_equality i1 i2 <-> extensional_equality i1 i2. +Proof. +Abort. + +(* ********** *) + +(* +Lemma isomorphism_equality_in_context_on_the_left : + forall x y i : Isomorphism, + x = y -> C x i = C y i. +Proof. +Abort. + +Proposition take_five : + forall i : Isomorphism, + extensional_equality (C (C (C (C i i) i) i) i) Id + -> + i = Id. +Proof. +Abort. + +Proposition characteristic_property_of_Id : + forall i : Isomorphism, + i = Id <-> forall v : Vertex, A i v = v. +Proof. +Abort. + +Proposition one_for_the_road : + forall i : Isomorphism, + (forall v : Vertex, A i v = v) + -> + C i i = Id. +Proof. +Abort. + +Proposition notable_property_of_Id : + exists i : Isomorphism, + exists v : Vertex, + not (A i v = v -> i = Id). +Proof. +Abort. + +Proposition notable_property_of_Id' : + not (forall (i : Isomorphism) (v : Vertex), + A i v = v -> i = Id). +Proof. +Abort. + +Proposition notable_property_of_symmetries : + forall (i : Isomorphism) + (v : Vertex), + A i v = v -> + ((i = IR R0) + \/ + (exists s : Reflection, + i = IS s)). +Proof. +Abort. + +Proposition notable_property_of_symmetries' : + forall i : Isomorphism, + (exists v : Vertex, + A i v = v) -> + ((i = IR R0) + \/ + (exists s : Reflection, + i = IS s)). +Proof. +Abort. + +Proposition one_more_for_the_road : + forall (i : Isomorphism) (v : Vertex), + A i v = v -> C i i = Id. +Proof. +Abort. +*) + +(* ********** *) + +(* end of week-07_isometries3.v *) diff --git a/cs3234/labs/week-07_soundness-and-completeness-of-equality-predicates-revisited.v b/cs3234/labs/week-07_soundness-and-completeness-of-equality-predicates-revisited.v new file mode 100644 index 0000000..b7da8f0 --- /dev/null +++ b/cs3234/labs/week-07_soundness-and-completeness-of-equality-predicates-revisited.v @@ -0,0 +1,281 @@ +(* week-07_soundness-and-completeness-of-equality-predicates-revisited.v *) +(* LPP 2024 - S3C234 2023-2024, Sem2 *) +(* Olivier Danvy *) +(* Version of 08 Mar 2024 *) + +(* ********** *) + +Ltac fold_unfold_tactic name := intros; unfold name; fold name; reflexivity. + +Require Import Arith Bool List. + +(* ********** *) + +Definition is_a_sound_and_complete_equality_predicate (V : Type) (eqb_V : V -> V -> bool) := + forall v1 v2 : V, + eqb_V v1 v2 = true <-> v1 = v2. + +(* ********** *) + +Check Bool.eqb. +(* eqb : bool -> bool -> bool *) + +Definition eqb_bool (b1 b2 : bool) : bool := + match b1 with + true => + match b2 with + true => + true + | false => + false + end + | false => + match b2 with + true => + false + | false => + true + end + end. + +Lemma eqb_bool_is_reflexive : + forall b : bool, + eqb_bool b b = true. +Proof. +Abort. + +Search (eqb _ _ = _ -> _ = _). +(* eqb_prop: forall a b : bool, eqb a b = true -> a = b *) + +Proposition soundness_and_completeness_of_eqb_bool : + is_a_sound_and_complete_equality_predicate bool eqb_bool. +Proof. + unfold is_a_sound_and_complete_equality_predicate. +Proof. +Abort. + +(* ***** *) + +Proposition soundness_and_completeness_of_eqb_bool : + is_a_sound_and_complete_equality_predicate bool eqb. +Proof. +Abort. + +(* ********** *) + +Check Nat.eqb. +(* Nat.eqb : nat -> nat -> bool *) + +Fixpoint eqb_nat (n1 n2 : nat) : bool := + match n1 with + O => + match n2 with + O => + true + | S n2' => + false + end + | S n1' => + match n2 with + O => + false + | S n2' => + eqb_nat n1' n2' + end + end. + +Lemma fold_unfold_eqb_nat_O : + forall n2 : nat, + eqb_nat 0 n2 = + match n2 with + O => + true + | S _ => + false + end. +Proof. + fold_unfold_tactic eqb_nat. +Qed. + +Lemma fold_unfold_eqb_nat_S : + forall n1' n2 : nat, + eqb_nat (S n1') n2 = + match n2 with + O => + false + | S n2' => + eqb_nat n1' n2' + end. +Proof. + fold_unfold_tactic eqb_nat. +Qed. + +Search (Nat.eqb _ _ = true -> _ = _). +(* beq_nat_true: forall n m : nat, (n =? m) = true -> n = m *) + +Proposition soundness_and_completeness_of_eqb_nat : + is_a_sound_and_complete_equality_predicate nat eqb_nat. +Proof. +Abort. + +(* ***** *) + +Lemma fold_unfold_eqb_Nat_O : + forall n2 : nat, + 0 =? n2 = + match n2 with + O => + true + | S _ => + false + end. +Proof. + fold_unfold_tactic Nat.eqb. +Qed. + +Lemma fold_unfold_eqb_Nat_S : + forall n1' n2 : nat, + S n1' =? n2 = + match n2 with + O => + false + | S n2' => + n1' =? n2' + end. +Proof. + fold_unfold_tactic Nat.eqb. +Qed. + +Proposition soundness_and_completeness_of_eqb_nat : + is_a_sound_and_complete_equality_predicate nat Nat.eqb. +Proof. +Abort. + +(* ********** *) + +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. + +Proposition soundness_and_completeness_of_eqb_option : + forall (V : Type) + (eqb_V : V -> V -> bool), + is_a_sound_and_complete_equality_predicate V eqb_V -> + is_a_sound_and_complete_equality_predicate (option V) (eqb_option V eqb_V). +Proof. +Abort. + +(* ***** *) + +(* +Definition eqb_option_option (V : Type) (eqb_V : V -> V -> bool) (oov1 oov2 : option (option V)) : bool := +*) + +(* +Proposition soundness_and_completeness_of_eqb_option_option : + forall (V : Type) + (eqb_V : V -> V -> bool), + is_a_sound_and_complete_equality_predicate V eqb_V -> + is_a_sound_and_complete_equality_predicate (option (option V)) (eqb_option_option V eqb_V). +Proof. +Abort. +*) + +(* ********** *) + +Definition eqb_pair (V : Type) (eqb_V : V -> V -> bool) (W : Type) (eqb_W : W -> W -> bool) (p1 p2 : V * W) : bool := + match p1 with + (v1, w1) => + match p2 with + (v2, w2) => + eqb_V v1 v2 && eqb_W w1 w2 + end + end. + +Proposition soundness_and_completeness_of_eqb_pair : + forall (V : Type) + (eqb_V : V -> V -> bool) + (W : Type) + (eqb_W : W -> W -> bool), + is_a_sound_and_complete_equality_predicate V eqb_V -> + is_a_sound_and_complete_equality_predicate W eqb_W -> + is_a_sound_and_complete_equality_predicate (V * W) (eqb_pair V eqb_V W eqb_W). +Proof. +Abort. + +(* ********** *) + +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. + +Proposition soundness_and_completeness_of_eqb_list : + forall (V : Type) + (eqb_V : V -> V -> bool), + is_a_sound_and_complete_equality_predicate V eqb_V -> + is_a_sound_and_complete_equality_predicate (list V) (eqb_list V eqb_V). +Proof. +Abort. + +(* ********** *) + +(* end of week-07_soundness-and-completeness-of-equality-predicates-revisited.v *) diff --git a/cs3234/labs/week-09_formalizing-two-by-two-matrices.v b/cs3234/labs/week-09_formalizing-two-by-two-matrices.v new file mode 100644 index 0000000..84c8a27 --- /dev/null +++ b/cs3234/labs/week-09_formalizing-two-by-two-matrices.v @@ -0,0 +1,126 @@ +(* week-09_formalizing-two-by-two-matrices.v *) +(* LPP 2024 - CS3234 2023-2024, Sem2 *) +(* Olivier Danvy *) +(* Version of Fri 22 Mar 2024 *) + +(* ********** *) + +Ltac fold_unfold_tactic name := intros; unfold name; fold name; reflexivity. + +Require Import Arith. + +(* ********** *) + +Inductive m22 : Type := M22 : nat -> nat -> nat -> nat -> m22. + +Property componential_equality_m22 : + forall x11 x12 x21 x22 y11 y12 y21 y22 : nat, + M22 x11 x12 + x21 x22 = + M22 y11 y12 + y21 y22 + <-> + x11 = y11 /\ x12 = y12 /\ x21 = y21 /\ x22 = y22. +Proof. + intros x11 x12 x21 x22 y11 y12 y21 y22. + split. + + - intro H_tmp. + injection H_tmp as H11 H12 H21 H22. + rewrite -> H11. + rewrite -> H12. + rewrite -> H21. + rewrite -> H22. + split; [reflexivity | split; [reflexivity | split; [reflexivity | reflexivity]]]. + + - intros [H11 [H12 [H21 H22]]]. + rewrite -> H11. + rewrite -> H12. + rewrite -> H21. + rewrite -> H22. + reflexivity. +Qed. + +(* ***** *) + +Definition zero_m22 := M22 0 0 + 0 0. + +Definition add_m22 (x y : m22) : m22 := + match (x, y) with + (M22 x11 x12 + x21 x22, + M22 y11 y12 + y21 y22) => + M22 (x11 + y11) (x12 + y12) + (x21 + y21) (x22 + y22) + end. + +Lemma add_m22_assoc : + forall x y z : m22, + add_m22 x (add_m22 y z) = + add_m22 (add_m22 x y) z. +Proof. + intros [x11 x12 + x21 x22] + [y11 y12 + y21 y22] + [z11 z12 + z21 z22]. + unfold add_m22. + rewrite ->4 Nat.add_assoc. + reflexivity. +Qed. + +Lemma add_m22_0_l : + forall x : m22, + add_m22 zero_m22 x = + x. +Proof. + intros [x11 x12 + x21 x22]. + unfold add_m22, zero_m22. + rewrite ->4 Nat.add_0_l. + reflexivity. +Qed. + +Lemma add_m22_0_r : + forall x : m22, + add_m22 x zero_m22 = + x. +Proof. + intros [x11 x12 + x21 x22]. + unfold add_m22, zero_m22. + rewrite ->4 Nat.add_0_r. + reflexivity. +Qed. + +(* ********** *) + +Inductive mm22 : Type := MM22 : m22 -> m22 -> m22 -> m22 -> mm22. + +Definition mul_m22 (x y : m22) := + match (x, y) with + (M22 x11 x12 + x21 x22, + M22 y11 y12 + y21 y22) => + M22 (x11 * y11 + x12 * y21) (x11 * y12 + x12 * y22) + (x21 * y11 + x22 * y21) (x21 * y12 + x22 * y22) + end. + +Property mul_m22_assoc : + forall (m1 m2 m3 : m22), + mul_m22 m1 (mul_m22 m2 m3) = (mul_m22 (mul_m22 m1 m2) m3). +Proof. + intros [m1_11 m1_12 m1_21 m1_22] + [m2_11 m2_12 m2_21 m2_22] + [m3_11 m3_12 m3_21 m3_22]. + unfold mul_m22. + rewrite -> 4 Nat.mul_add_distr_r. +Qed. + +(* ********** *) + +(* week-09_formalizing-two-by-two-matrices.v *) diff --git a/cs3234/labs/week-10_a-fibonacci-structure.v b/cs3234/labs/week-10_a-fibonacci-structure.v new file mode 100644 index 0000000..104670c --- /dev/null +++ b/cs3234/labs/week-10_a-fibonacci-structure.v @@ -0,0 +1,1311 @@ +(* week-10_a-fibonacci-structure.v *) +(* LPP 2024 - CS3233 2023-2024, Sem2 *) +(* Olivier Danvy *) +(* Version of 04 Apr 2024 *) + +(* ********** *) + +Ltac fold_unfold_tactic name := intros; unfold name; fold name; reflexivity. + +Require Import Arith. + +(* ********** *) + +(* The Fibonacci structure: *) + +Definition Fibonacci (career : Type) + (add : career -> career -> career) + (fib_0 : career) + (fib_1 : career) + (fib : nat -> career) := + (fib 0 = fib_0) + /\ + (fib 1 = fib_1) + /\ + (forall n : nat, + fib (S (S n)) = add (fib n) (fib (S n))). + +(* ********** *) + +(* The standard Fibonacci function over nats: *) + +Fixpoint fib_nat (n : nat) : nat := + match n with + 0 => 0 + | S n' => match n' with + 0 => 1 + | S n'' => fib_nat n'' + fib_nat n' + end + end. + +(* The standard fold-unfold lemmas: *) + +Lemma fold_unfold_fib_nat_O : + fib_nat 0 = + 0. +Proof. + fold_unfold_tactic fib_nat. +Qed. + +Lemma fold_unfold_fib_nat_S : + forall n' : nat, + fib_nat (S n') = + match n' with + 0 => 1 + | S n'' => fib_nat n'' + fib_nat n' + end. +Proof. + fold_unfold_tactic fib_nat. +Qed. + +(* Two dedicated fold-unfold lemmas: *) + +Lemma fold_unfold_fib_nat_1 : + fib_nat 1 = + 1. +Proof. + exact (fold_unfold_fib_nat_S 0). +Qed. + +Lemma fold_unfold_fib_nat_SS : + forall n'' : nat, + fib_nat (S (S n'')) = + fib_nat n'' + fib_nat (S n''). +Proof. + intro n''. + exact (fold_unfold_fib_nat_S (S n'')). +Qed. + +(* The standard Fibonacci function over nats fits the Fibonacci structure: *) + +Theorem fib_nat_fits_Fibonacci : + Fibonacci nat Nat.add 0 1 fib_nat. +Proof. + unfold Fibonacci. + split. + exact fold_unfold_fib_nat_O. + split. + exact fold_unfold_fib_nat_1. + exact fold_unfold_fib_nat_SS. +Qed. + +(* ********** *) + +(* Computing Fibonacci numbers by exponentiating a 2x2 matrix of nats: *) + +Inductive m22_nat : Type := + | M22_nat : nat -> nat -> nat -> nat -> m22_nat. + +Property componential_equality_m22_nat : + forall x11 x12 x21 x22 y11 y12 y21 y22 : nat, + M22_nat x11 x12 + x21 x22 = + M22_nat y11 y12 + y21 y22 + <-> + x11 = y11 /\ x12 = y12 /\ x21 = y21 /\ x22 = y22. +Proof. + intros x11 x12 x21 x22 y11 y12 y21 y22. + split. + + - intro H_tmp. + injection H_tmp as H11 H12 H21 H22. + rewrite -> H11. + rewrite -> H12. + rewrite -> H21. + rewrite -> H22. + split; [reflexivity | split; [reflexivity | split; [reflexivity | reflexivity]]]. + + - intros [H11 [H12 [H21 H22]]]. + rewrite -> H11. + rewrite -> H12. + rewrite -> H21. + rewrite -> H22. + reflexivity. +Qed. + +Definition zero_m22_nat := M22_nat 0 0 + 0 0. + +Definition add_m22_nat (x y : m22_nat) : m22_nat := + match (x, y) with + (M22_nat x11 x12 + x21 x22, + M22_nat y11 y12 + y21 y22) => + M22_nat (x11 + y11) (x12 + y12) + (x21 + y21) (x22 + y22) + end. + +Lemma add_m22_nat_assoc : + forall x y z : m22_nat, + add_m22_nat x (add_m22_nat y z) = + add_m22_nat (add_m22_nat x y) z. +Proof. + intros [x11 x12 + x21 x22] + [y11 y12 + y21 y22] + [z11 z12 + z21 z22]. + unfold add_m22_nat. + rewrite ->4 Nat.add_assoc. + reflexivity. +Qed. + +Lemma add_m22_nat_0_l : + forall x : m22_nat, + add_m22_nat zero_m22_nat x = + x. +Proof. + intros [x11 x12 + x21 x22]. + unfold add_m22_nat, zero_m22_nat. + rewrite ->4 Nat.add_0_l. + reflexivity. +Qed. + +Lemma add_m22_nat_0_r : + forall x : m22_nat, + add_m22_nat x zero_m22_nat = + x. +Proof. + intros [x11 x12 + x21 x22]. + unfold add_m22_nat, zero_m22_nat. + rewrite ->4 Nat.add_0_r. + reflexivity. +Qed. + +Definition one_m22_nat := M22_nat 1 0 + 0 1. + +Definition mul_m22_nat (x y : m22_nat) : m22_nat := + match (x, y) with + (M22_nat x11 x12 + x21 x22, + M22_nat y11 y12 + y21 y22) => + (M22_nat (x11 * y11 + x12 * y21) (x11 * y12 + x12 * y22) + (x21 * y11 + x22 * y21) (x21 * y12 + x22 * y22)) + end. + +Lemma mul_m22_nat_assoc : + forall x y z : m22_nat, + mul_m22_nat x (mul_m22_nat y z) = + mul_m22_nat (mul_m22_nat x y) z. +Proof. +Admitted. + +Lemma mul_m22_nat_1_l : + forall x : m22_nat, + mul_m22_nat one_m22_nat x = x. +Proof. + intros [x11 x12 + x21 x22]. + unfold mul_m22_nat, one_m22_nat. + rewrite ->4 Nat.mul_0_l. + rewrite ->2 Nat.add_0_l. + rewrite ->2 Nat.add_0_r. + rewrite ->4 Nat.mul_1_l. + reflexivity. +Qed. + +Lemma mul_m22_nat_1_r : + forall x : m22_nat, + mul_m22_nat x one_m22_nat = + x. +Proof. + intros [x11 x12 + x21 x22]. + unfold mul_m22_nat, one_m22_nat. + rewrite ->4 Nat.mul_0_r. + rewrite ->2 Nat.add_0_r. + rewrite ->2 Nat.add_0_l. + rewrite ->4 Nat.mul_1_r. + reflexivity. +Qed. + +Lemma mul_m22_nat_0_l : + forall x : m22_nat, + mul_m22_nat zero_m22_nat x = + zero_m22_nat. +Proof. + intros [x11 x12 + x21 x22]. + unfold mul_m22_nat, zero_m22_nat. + rewrite ->4 Nat.mul_0_l. + rewrite -> Nat.add_0_l. + reflexivity. +Qed. + +Lemma mul_m22_0_r : + forall x : m22_nat, + mul_m22_nat x zero_m22_nat = + zero_m22_nat. +Proof. + intros [x11 x12 + x21 x22]. + unfold mul_m22_nat, zero_m22_nat. + rewrite ->4 Nat.mul_0_r. + rewrite -> Nat.add_0_r. + reflexivity. +Qed. + +Fixpoint power_m22_nat (x : m22_nat) (n : nat) : m22_nat:= + match n with + 0 => + one_m22_nat + | S n' => + mul_m22_nat x (power_m22_nat x n') + end. + +Lemma fold_unfold_power_m22_nat_O : + forall x : m22_nat, + power_m22_nat x 0 = + one_m22_nat. +Proof. + fold_unfold_tactic power_m22_nat. +Qed. + +Lemma fold_unfold_power_m22_nat_S : + forall (x : m22_nat) + (n' : nat), + power_m22_nat x (S n') = + mul_m22_nat x (power_m22_nat x n'). +Proof. + fold_unfold_tactic power_m22_nat. +Qed. + +Property about_exponentiating_2x2_matrices_of_nats : + forall (x : m22_nat) + (p q : nat), + power_m22_nat x (p + q) = + mul_m22_nat (power_m22_nat x p) (power_m22_nat x q). +Proof. + intros x p. + induction p as [ | p' IHp']; intro q. + + - rewrite -> Nat.add_0_l. + rewrite -> fold_unfold_power_m22_nat_O. + rewrite -> mul_m22_nat_1_l. + reflexivity. + + - rewrite -> plus_Sn_m. + rewrite -> fold_unfold_power_m22_nat_S. + rewrite -> (IHp' q). + rewrite -> fold_unfold_power_m22_nat_S. + apply mul_m22_nat_assoc. +Qed. + +(* The Fibonacci matrix: *) + +Definition F_nat := M22_nat 0 1 + 1 1. + +Definition F_nat_0 := power_m22_nat F_nat 0. +Compute F_nat_0. (* = M22_nat 1 0 0 1 *) +Definition F_nat_1 := power_m22_nat F_nat 1. +Compute F_nat_1. (* = M22_nat 0 1 1 1 *) +Definition F_nat_2 := power_m22_nat F_nat 2. +Compute F_nat_2. (* = M22_nat 1 1 1 2 *) +Definition F_nat_3 := power_m22_nat F_nat 3. +Compute F_nat_3. (* = M22_nat 1 2 2 3 *) +Definition F_nat_4 := power_m22_nat F_nat 4. +Compute F_nat_4. (* = M22_nat 2 3 3 5 *) + +Proposition about_exponentiating_F_nat : + forall n : nat, + power_m22_nat F_nat (S n) = + M22_nat (fib_nat n) (fib_nat (S n)) + (fib_nat (S n)) (fib_nat (S (S n))). +Proof. + intro n. + induction n as [ | n' IHn']. + + - rewrite -> fold_unfold_power_m22_nat_S. + rewrite -> fold_unfold_power_m22_nat_O. + rewrite -> mul_m22_nat_1_r. + unfold F_nat. + rewrite -> fold_unfold_fib_nat_SS. + rewrite -> fold_unfold_fib_nat_O. + rewrite -> fold_unfold_fib_nat_1. + rewrite -> Nat.add_0_l. + reflexivity. + + - rewrite -> fold_unfold_power_m22_nat_S. + rewrite -> IHn'. + unfold F_nat. + unfold mul_m22_nat. + rewrite -> Nat.mul_0_l. + rewrite -> Nat.add_0_l. + rewrite -> Nat.mul_1_l. + rewrite -> Nat.mul_0_l. + rewrite -> Nat.add_0_l. + rewrite -> Nat.mul_1_l. + rewrite -> Nat.mul_1_l. + rewrite <- fold_unfold_fib_nat_SS. + rewrite <- fold_unfold_fib_nat_SS. + reflexivity. +Qed. + +(* The following property is a corollary of + Property about_exponentiating_2x2_matrices_of_nats + and + Proposition about_exponentiating_F_nat + about the first cell of + mult_m22 (power_m22_nat F_nat (S p)) (power_m22_nat F_nat (S q)) + for any nat p and q: +*) + +Corollary an_identity_about_Fibonacci_numbers : + forall p q : nat, + fib_nat (S (p + q)) = + fib_nat p * fib_nat q + fib_nat (S p) * fib_nat (S q). +Proof. +Abort. + +(* ********** *) + +(* A perhaps unexpected property: *) + +Property as_luck_has_it : + forall n : nat, + power_m22_nat F_nat (S (S n)) = + add_m22_nat (power_m22_nat F_nat n) (power_m22_nat F_nat (S n)). +Proof. +Admitted. + +Definition fib_m22_nat (n : nat) : m22_nat := + power_m22_nat F_nat n. + +Compute fib_m22_nat 0. (* = M22_nat 1 0 0 1 *) +Compute fib_m22_nat 1. (* = M22_nat 0 1 1 1 *) +Compute fib_m22_nat 2. (* = M22_nat 1 1 1 2 *) +Compute fib_m22_nat 3. (* = M22_nat 1 2 2 3 *) + +(* Let's make our own luck with fortuitous fold_unfold_lemmas: *) + +Lemma fold_unfold_fib_m22_nat_0 : + fib_m22_nat 0 = + one_m22_nat. +Proof. + unfold fib_m22_nat. + apply fold_unfold_power_m22_nat_O. +Qed. + +Lemma fold_unfold_fib_m22_nat_1 : + fib_m22_nat 1 = + F_nat. +Proof. + unfold fib_m22_nat. + rewrite -> fold_unfold_power_m22_nat_S. + rewrite -> fold_unfold_power_m22_nat_O. + apply mul_m22_nat_1_r. +Qed. + +Lemma fold_unfold_fib_m22_nat_SS : + forall (n : nat), + fib_m22_nat (S (S n)) = + add_m22_nat (fib_m22_nat n) (fib_m22_nat (S n)). +Proof. + intro n. + unfold fib_m22_nat. + apply as_luck_has_it. +Qed. + +(* So successive exponents of F_nat fit the Fibonacci structure: *) + +Theorem fib_m22_nat_fits_Fibonacci : + Fibonacci m22_nat add_m22_nat one_m22_nat F_nat fib_m22_nat. +Proof. + unfold Fibonacci. + split. + exact fold_unfold_fib_m22_nat_0. + split. + exact fold_unfold_fib_m22_nat_1. + exact fold_unfold_fib_m22_nat_SS. +Qed. + +(* Can this property be generalized? *) + +(* ********** *) + +(* Polymorphic 2x2 matrices: *) + +Inductive pm22 (V : Type) := + | PM22 : V -> V -> V -> V -> pm22 V. + +(* ***** *) + +(* For example, 2x2 matrices of nats: *) + +Check (pm22 nat). + +(* For example, 2x2 matrices of 2x2 matrices of nats: *) + +Check (pm22 (pm22 nat)). + +(* ***** *) + +Definition zero_pm22 (V : Type) (zero : V) : pm22 V := + PM22 V + zero zero + zero zero. + +Definition add_pm22 (V : Type) (add : V -> V -> V) (x y : pm22 V) : pm22 V := + match (x, y) with + (PM22 _ + x11 x12 + x21 x22, + PM22 _ + y11 y12 + y21 y22) => + PM22 V + (add x11 y11) (add x12 y12) + (add x21 y21) (add x22 y22) + end. + +Property add_pm22_assoc : + forall (V : Type) + (add : V -> V -> V), + (forall x y z : V, + add x (add y z) = add (add x y) z) -> + forall x y z : pm22 V, + add_pm22 V add x (add_pm22 V add y z) = + add_pm22 V add (add_pm22 V add x y) z. +Proof. + intros V add H_add_assoc + [x11 x12 + x21 x22] + [y11 y12 + y21 y22] + [z11 z12 + z21 z22]. + unfold add_pm22. + rewrite ->4 H_add_assoc. + reflexivity. +Qed. + +Property add_pm22_0_l : + forall (V : Type) + (add : V -> V -> V) + (zero : V), + (forall v : V, + add zero v = v) -> + forall x : pm22 V, + add_pm22 V add (zero_pm22 V zero) x = + x. +Proof. + intros V add zero H_add_0_l + [x11 x12 + x21 x22]. + unfold add_pm22, zero_pm22. + rewrite ->4 H_add_0_l. + reflexivity. +Qed. + +Property add_pm22_0_r : + forall (V : Type) + (add : V -> V -> V) + (zero : V), + (forall v : V, + add v zero = v) -> + forall x : pm22 V, + add_pm22 V add x (zero_pm22 V zero) = + x. +Proof. + intros V add zero H_add_0_r + [x11 x12 + x21 x22]. + unfold add_pm22, zero_pm22. + rewrite ->4 H_add_0_r. + reflexivity. +Qed. + +Definition one_pm22 (V : Type) (zero : V) (one : V) : pm22 V := + PM22 V + one zero + zero one. + +Definition mul_pm22 (V : Type) (add : V -> V -> V) (mul : V -> V -> V) (x y : pm22 V) : pm22 V := + match (x, y) with + (PM22 _ + x11 x12 + x21 x22, + PM22 _ + y11 y12 + y21 y22) => + PM22 V + (add (mul x11 y11) (mul x12 y21)) (add (mul x11 y12) (mul x12 y22)) + (add (mul x21 y11) (mul x22 y21)) (add (mul x21 y12) (mul x22 y22)) + end. + +Property mul_pm22_assoc : + forall (V : Type) + (add : V -> V -> V), + (forall x y z : V, + add x (add y z) = add (add x y) z) -> + (forall x y : V, + add x y = add y x) -> + forall mul : V -> V -> V, + (forall x y z : V, + mul x (mul y z) = mul (mul x y) z) -> + (forall x y z : V, + mul x (add y z) = add (mul x y) (mul x z)) -> + (forall x y z : V, + mul (add x y) z = add (mul x z) (mul y z)) -> + forall x y z : pm22 V, + mul_pm22 V add mul x (mul_pm22 V add mul y z) = + mul_pm22 V add mul (mul_pm22 V add mul x y) z. +Proof. +Admitted. + +Property mul_pm22_1_l : + forall (V : Type) + (add : V -> V -> V) + (mul : V -> V -> V) + (zero : V), + (forall v : V, + add zero v = v) -> + (forall v : V, + add v zero = v) -> + (forall v : V, + mul zero v = zero) -> + forall (one : V), + (forall v : V, + mul one v = v) -> + forall x : pm22 V, + mul_pm22 V add mul (one_pm22 V zero one) x = + x. +Proof. + intros V add mul zero H_add_0_l H_add_0_r H_mul_0_l one H_mul_1_l + [x11 x12 + x21 x22]. + unfold mul_pm22, one_pm22. + rewrite ->4 H_mul_0_l. + rewrite ->2 H_add_0_l. + rewrite ->2 H_add_0_r. + rewrite ->4 H_mul_1_l. + reflexivity. +Qed. + +Property mul_pm22_1_r : + forall (V : Type) + (add : V -> V -> V) + (mul : V -> V -> V) + (zero : V), + (forall v : V, + add zero v = v) -> + (forall v : V, + add v zero = v) -> + (forall v : V, + mul v zero = zero) -> + forall (one : V), + (forall v : V, + mul v one = v) -> + forall x : pm22 V, + mul_pm22 V add mul x (one_pm22 V zero one) = + x. +Proof. + intros V add mul zero H_add_0_l H_add_0_r H_mul_0_r one H_mul_1_r + [x11 x12 + x21 x22]. + unfold mul_pm22, one_pm22. + rewrite ->4 H_mul_0_r. + rewrite ->2 H_add_0_l. + rewrite ->2 H_add_0_r. + rewrite ->4 H_mul_1_r. + reflexivity. +Qed. + +Property mul_pm22_0_l : + forall (V : Type) + (add : V -> V -> V) + (mul : V -> V -> V) + (zero : V), + (forall v : V, + add zero v = v) -> + (forall v : V, + mul zero v = zero) -> + forall x : pm22 V, + mul_pm22 V add mul (zero_pm22 V zero) x = + zero_pm22 V zero. +Proof. + intros V add mul zero H_add_0_l H_mul_0_l + [x11 x12 + x21 x22]. + unfold mul_pm22, zero_pm22. + rewrite ->4 H_mul_0_l. + rewrite -> H_add_0_l. + reflexivity. +Qed. + +Property mul_pm22_0_r : + forall (V : Type) + (add : V -> V -> V) + (mul : V -> V -> V) + (zero : V), + (forall v : V, + add v zero = v) -> + (forall v : V, + mul v zero = zero) -> + forall x : pm22 V, + mul_pm22 V add mul x (zero_pm22 V zero) = + zero_pm22 V zero. +Proof. + intros V add mul zero H_add_0_r H_mul_0_r + [x11 x12 + x21 x22]. + unfold mul_pm22, zero_pm22. + rewrite ->4 H_mul_0_r. + rewrite -> H_add_0_r. + reflexivity. +Qed. + +(* power_pm22 : forall V : Type, (V -> V -> V) -> (V -> V -> V) -> V -> V -> pm22 V -> nat -> pm22 V *) +Fixpoint power_pm22 (V : Type) (add mul : V -> V -> V) (zero one : V) (x : pm22 V) (n : nat) : pm22 V := + match n with + 0 => + one_pm22 V zero one + | S n' => + mul_pm22 V add mul x (power_pm22 V add mul zero one x n') + end. + +Lemma fold_unfold_power_pm22_O : + forall (V : Type) + (add mul : V -> V -> V) + (zero one : V) + (x : pm22 V), + power_pm22 V add mul zero one x 0 = + one_pm22 V zero one. +Proof. + fold_unfold_tactic power_pm22. +Qed. + +Lemma fold_unfold_power_pm22_S : + forall (V : Type) + (add mul : V -> V -> V) + (zero one : V) + (x : pm22 V) + (n' : nat), + power_pm22 V add mul zero one x (S n') = + mul_pm22 V add mul x (power_pm22 V add mul zero one x n'). +Proof. + fold_unfold_tactic power_pm22. +Qed. + +Definition F_pm22 (V : Type) (zero : V) (one : V) : pm22 V := + PM22 V + zero one + one one. + +Definition fib_pm22 (V : Type) (add mul : V -> V -> V) (zero one : V) (n : nat) : pm22 V := + power_pm22 V add mul zero one (F_pm22 V zero one) n. + +Lemma fold_unfold_fib_pm22_0 : + forall (V : Type) + (add mul : V -> V -> V) + (zero one : V), + fib_pm22 V add mul zero one 0 = + one_pm22 V zero one. +Proof. + intros V add mul zero one. + unfold fib_pm22. + exact (fold_unfold_power_pm22_O V add mul zero one (F_pm22 V zero one)). +Qed. + +Lemma fold_unfold_fib_pm22_1 : + forall (V : Type) + (add mul : V -> V -> V) + (zero : V), + (forall v : V, + add zero v = v) -> + (forall v : V, + add v zero = v) -> + (forall v : V, + mul v zero = zero) -> + forall (one : V), + (forall v : V, + mul v one = v) -> + fib_pm22 V add mul zero one 1 = + F_pm22 V zero one. +Proof. + intros V add mul zero H_add_0_l H_add_0_r H_mul_0_r one H_mul_1_r. + unfold fib_pm22. + rewrite -> fold_unfold_power_pm22_S. + rewrite -> fold_unfold_power_pm22_O. + rewrite -> (mul_pm22_1_r V add mul zero H_add_0_l H_add_0_r H_mul_0_r one H_mul_1_r). + reflexivity. +Qed. + +Lemma fold_unfold_fib_pm22_SS : + forall (V : Type) + (add mul : V -> V -> V) + (zero : V), + (forall v : V, + add zero v = v) -> + (forall v : V, + mul zero v = zero) -> + forall (one : V), + (forall v : V, + mul one v = v) -> + forall (n : nat), + fib_pm22 V add mul zero one (S (S n)) = + add_pm22 V + add + (fib_pm22 V add mul zero one n) + (fib_pm22 V add mul zero one (S n)). +Proof. + intros V add mul zero H_add_0_l H_mul_0_l one H_mul_1_l n. + unfold fib_pm22. + rewrite -> fold_unfold_power_pm22_S. + rewrite -> fold_unfold_power_pm22_S. + destruct (power_pm22 V add mul zero one (F_pm22 V zero one) n) as [x11 x12 + x21 x22] eqn:H_the_guy. + unfold F_pm22. + unfold mul_pm22, add_pm22. + repeat (rewrite -> H_mul_0_l). + repeat (rewrite -> H_mul_1_l). + repeat (rewrite -> H_add_0_l). + reflexivity. +Qed. + +Theorem fib_pm22_fits_Fibonacci : + forall (V : Type) + (zero one : V) + (add mul : V -> V -> V), + (forall v : V, + add zero v = v) -> + (forall v : V, + add v zero = v) -> + (forall v : V, + mul zero v = zero) -> + (forall v : V, + mul v zero = zero) -> + (forall v : V, + mul one v = v) -> + (forall v : V, + mul v one = v) -> + Fibonacci (pm22 V) + (add_pm22 V add) + (one_pm22 V zero one) + (F_pm22 V zero one) + (fib_pm22 V add mul zero one). +Proof. + intros V zero one add mul add_0_l add_0_r mul_0_l mul_0_r mul_1_l mul_1_r. + unfold Fibonacci. + split. + exact (fold_unfold_fib_pm22_0 V add mul zero one). + split. + exact (fold_unfold_fib_pm22_1 V add mul zero add_0_l add_0_r mul_0_r one mul_1_r). + exact (fold_unfold_fib_pm22_SS V add mul zero add_0_l mul_0_l one mul_1_l). +Qed. + +(* ********** *) + +(* Revisiting 2x2 matrices of nats: *) + +Definition pm22_nat := pm22 nat. + +Definition zero_pm22_nat := zero_pm22 nat 0. + +Definition add_pm22_nat := add_pm22 nat Nat.add. + +Lemma add_pm22_nat_assoc : + forall x y z : pm22_nat, + add_pm22_nat x (add_pm22_nat y z) = + add_pm22_nat (add_pm22_nat x y) z. +Proof. + exact (add_pm22_assoc nat Nat.add Nat.add_assoc). +Qed. + +Lemma add_pm22_nat_0_l : + forall x : pm22_nat, + add_pm22_nat zero_pm22_nat x = + x. +Proof. + exact (add_pm22_0_l nat Nat.add 0 Nat.add_0_l). +Qed. + +Lemma add_pm22_nat_0_r : + forall x : pm22_nat, + add_pm22_nat x zero_pm22_nat = + x. +Proof. + exact (add_pm22_0_r nat Nat.add 0 Nat.add_0_r). +Qed. + +Definition one_pm22_nat := one_pm22 nat 0 1. + +Definition mul_pm22_nat := mul_pm22 nat Nat.add Nat.mul. + +Lemma mul_pm22_nat_assoc : + forall x y z : pm22_nat, + mul_pm22_nat x (mul_pm22_nat y z) = + mul_pm22_nat (mul_pm22_nat x y) z. +Proof. + exact (mul_pm22_assoc nat Nat.add Nat.add_assoc Nat.add_comm Nat.mul Nat.mul_assoc Nat.mul_add_distr_l Nat.mul_add_distr_r). +Qed. + +Lemma mul_pm22_nat_1_l : + forall x : pm22_nat, + mul_pm22_nat one_pm22_nat x = + x. +Proof. + exact (mul_pm22_1_l nat Nat.add Nat.mul 0 Nat.add_0_l Nat.add_0_r Nat.mul_0_l 1 Nat.mul_1_l). +Qed. + +Lemma mul_pm22_nat_1_r : + forall x : pm22_nat, + mul_pm22_nat x one_pm22_nat = + x. +Proof. + exact (mul_pm22_1_r nat Nat.add Nat.mul 0 Nat.add_0_l Nat.add_0_r Nat.mul_0_r 1 Nat.mul_1_r). +Qed. + +Lemma mul_pm22_nat_0_l : + forall x : pm22_nat, + mul_pm22_nat zero_pm22_nat x = + zero_pm22_nat. +Proof. + exact (mul_pm22_0_l nat Nat.add Nat.mul 0 Nat.add_0_l Nat.mul_0_l). +Qed. + +Lemma mul_pm22_nat_0_r : + forall x : pm22_nat, + mul_pm22_nat x zero_pm22_nat = + zero_pm22_nat. +Proof. + exact (mul_pm22_0_r nat Nat.add Nat.mul 0 Nat.add_0_r Nat.mul_0_r). +Qed. + +Definition power_pm22_nat := power_pm22 nat Nat.add Nat.mul 0 1. + +Lemma fold_unfold_power_pm22_nat_O : + forall x : pm22_nat, + power_pm22_nat x 0 = + one_pm22_nat. +Proof. + exact (fold_unfold_power_pm22_O nat Nat.add Nat.mul 0 1). +Qed. + +Lemma fold_unfold_power_pm22_nat_S : + forall (x : pm22_nat) + (n' : nat), + power_pm22_nat x (S n') = + mul_pm22_nat x (power_pm22_nat x n'). +Proof. + exact (fold_unfold_power_pm22_S nat Nat.add Nat.mul 0 1). +Qed. + +Definition F_pm22_nat := F_pm22 nat 0 1. + +(* +Definition fib_pm22_nat (n : nat) : pm22_nat := + power_pm22_nat F_pm22_nat n. +*) + +Definition fib_pm22_nat := fib_pm22 nat Nat.add Nat.mul 0 1. + +Lemma fold_unfold_fib_pm22_nat_0 : + fib_pm22_nat 0 = + one_pm22_nat. +Proof. + exact (fold_unfold_fib_pm22_0 nat Nat.add Nat.mul 0 1). +Qed. + +Lemma fold_unfold_fib_pm22_nat_1 : + fib_pm22_nat 1 = + F_pm22_nat. +Proof. + exact (fold_unfold_fib_pm22_1 nat Nat.add Nat.mul 0 Nat.add_0_l Nat.add_0_r Nat.mul_0_r 1 Nat.mul_1_r). +Qed. + +Lemma fold_unfold_fib_pm22_nat_SS : + forall n : nat, + fib_pm22_nat (S (S n)) = + add_pm22_nat (fib_pm22_nat n) (fib_pm22_nat (S n)). +Proof. + exact (fold_unfold_fib_pm22_SS nat Nat.add Nat.mul 0 Nat.add_0_l Nat.mul_0_l 1 Nat.mul_1_l). +Qed. + +Corollary fib_pm22_nat_fits_Fibonacci : + Fibonacci pm22_nat add_pm22_nat one_pm22_nat F_pm22_nat fib_pm22_nat. +Proof. + exact (fib_pm22_fits_Fibonacci nat 0 1 Nat.add Nat.mul Nat.add_0_l Nat.add_0_r Nat.mul_0_l Nat.mul_0_r Nat.mul_1_l Nat.mul_1_r). +Qed. + +(* ********** *) + +Theorem about_layering : + forall (V : Type) + (zero one : V) + (add mul : V -> V -> V), + (forall v : V, + add zero v = v) -> + (forall v : V, + add v zero = v) -> + (forall v : V, + mul zero v = zero) -> + (forall v : V, + mul v zero = zero) -> + (forall v : V, + mul one v = v) -> + (forall v : V, + mul v one = v) -> + forall fib : nat -> V, + Fibonacci V add zero one fib -> + Fibonacci (pm22 V) + (add_pm22 V add) + (one_pm22 V (fib 0) (fib 1)) + (F_pm22 V (fib 0) (fib 1)) + (fib_pm22 V add mul (fib 0) (fib 1)). +Proof. + intros V zero one add mul add_0_l add_0_r mul_0_l mul_0_r mul_1_l mul_1_r fib. + unfold Fibonacci. + intros [H_fib_0 [H_fib_1 H_fib_SS]]. + split. + - rewrite -> fold_unfold_fib_pm22_0. + reflexivity. + - split. + + rewrite -> H_fib_0. + rewrite -> H_fib_1. + Check (fold_unfold_fib_pm22_1 V add mul zero add_0_l add_0_r mul_0_r one mul_1_r). + exact (fold_unfold_fib_pm22_1 V add mul zero add_0_l add_0_r mul_0_r one mul_1_r). + + rewrite -> H_fib_0. + rewrite -> H_fib_1. + Check (fold_unfold_fib_pm22_SS V add mul zero add_0_l mul_0_l one mul_1_l). + exact (fold_unfold_fib_pm22_SS V add mul zero add_0_l mul_0_l one mul_1_l). +Qed. + +(* ********** *) + +Proposition about_fib_pm22_nat : + forall n : nat, + fib_pm22_nat (S n) = + PM22 nat + (fib_nat n) (fib_nat (S n)) + (fib_nat (S n)) (fib_nat (S (S n))). +Proof. + intro n. + unfold fib_pm22_nat. + unfold fib_pm22. + induction n as [ | n' IHn']. + + - rewrite -> fold_unfold_power_pm22_nat_S. + rewrite -> fold_unfold_power_pm22_nat_O. + rewrite -> mul_pm22_nat_1_r. + unfold F_pm22_nat. + unfold F_pm22. + rewrite ->3 fold_unfold_fib_nat_S. + rewrite -> fold_unfold_fib_nat_O. + rewrite -> Nat.add_0_l. + reflexivity. + + - rewrite -> fold_unfold_power_pm22_nat_S. + unfold power_pm22_nat. + rewrite -> IHn'. + unfold mul_pm22_nat, mul_pm22, F_pm22_nat, F_pm22. + rewrite -> Nat.mul_0_l. + rewrite -> Nat.mul_1_l. + rewrite -> Nat.add_0_l. + rewrite -> Nat.mul_0_l. + rewrite -> Nat.mul_1_l. + rewrite -> Nat.add_0_l. + rewrite -> Nat.mul_1_l. + rewrite <-2 fold_unfold_fib_nat_SS. + reflexivity. +Qed. + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (zero_pm22 nat 0) (F_pm22 nat 0 1) 0). +(* = PM22 (pm22 nat) (PM22 nat 0 1 1 1) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 0 1 1 1) *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (zero_pm22 nat 0) (F_pm22 nat 0 1) 1). +(* = PM22 (pm22 nat) (PM22 nat 0 0 0 0) (PM22 nat 1 1 1 2) (PM22 nat 1 1 1 2) (PM22 nat 1 1 1 2) *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (zero_pm22 nat 0) (F_pm22 nat 0 1) 2). +(* = PM22 (pm22 nat) (PM22 nat 1 2 2 3) (PM22 nat 1 2 2 3) (PM22 nat 1 2 2 3) (PM22 nat 2 4 4 6) *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (zero_pm22 nat 0) (F_pm22 nat 0 1) 3). +(* = PM22 (pm22 nat) (PM22 nat 2 3 3 5) (PM22 nat 4 6 6 10) (PM22 nat 4 6 6 10) (PM22 nat 6 9 9 15) *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (one_pm22 nat 0 1) (F_pm22 nat 0 1) 0). +(* = PM22 (pm22 nat) (PM22 nat 0 1 1 1) (PM22 nat 1 0 0 1) (PM22 nat 1 0 0 1) (PM22 nat 0 1 1 1) *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (one_pm22 nat 0 1) (F_pm22 nat 0 1) 1). +(* = PM22 (pm22 nat) (PM22 nat 0 2 2 2) (PM22 nat 2 1 1 3) (PM22 nat 1 2 2 3) (PM22 nat 1 2 2 3) *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (one_pm22 nat 0 1) (F_pm22 nat 0 1) 2). +(* = PM22 (pm22 nat) (PM22 nat 2 5 5 7) (PM22 nat 4 4 4 8) (PM22 nat 4 5 5 9) (PM22 nat 3 6 6 9) *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (one_pm22 nat 0 1) (F_pm22 nat 0 1) 3). +(* = PM22 (pm22 nat) (PM22 nat 7 14 14 21) (PM22 nat 10 13 13 23) (PM22 nat 10 16 16 26) (PM22 nat 10 17 17 27) *) + +Property as_luck_has_it_too : + forall n : nat, + fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (zero_pm22 nat 0) (one_pm22 nat 0 1) (S n) = + PM22 (pm22 nat) + (PM22 nat (fib_nat n) 0 0 (fib_nat n)) + (PM22 nat (fib_nat (S n)) 0 0 (fib_nat (S n))) + (PM22 nat (fib_nat (S n)) 0 0 (fib_nat (S n))) + (PM22 nat (fib_nat (S (S n))) 0 0 (fib_nat (S (S n)))). +Proof. + intro n. + induction n as [ | n' IHn']. + + - rewrite -> (fold_unfold_fib_pm22_1 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (zero_pm22 nat 0) (add_pm22_0_l nat Nat.add 0 Nat.add_0_l) (add_pm22_0_r nat Nat.add 0 Nat.add_0_r) (mul_pm22_0_r nat Nat.add Nat.mul 0 Nat.add_0_r Nat.mul_0_r) (one_pm22 nat 0 1) (mul_pm22_1_r nat Nat.add Nat.mul 0 Nat.add_0_l Nat.add_0_r Nat.mul_0_r 1 Nat.mul_1_r)). + unfold one_pm22. + unfold zero_pm22. + unfold F_pm22. + repeat (rewrite -> fold_unfold_fib_nat_SS). + repeat (rewrite -> fold_unfold_fib_nat_1). + repeat (rewrite -> fold_unfold_fib_nat_O). + repeat (rewrite -> Nat.add_0_l). + reflexivity. + + - unfold fib_pm22. + unfold fib_pm22 in IHn'. + rewrite -> fold_unfold_power_pm22_S. + rewrite -> IHn'. + unfold F_pm22. + unfold zero_pm22. + unfold one_pm22. + unfold mul_pm22. + unfold add_pm22. + repeat (rewrite -> Nat.add_0_l). + repeat (rewrite -> Nat.mul_0_l). + repeat (rewrite -> Nat.mul_1_l). + repeat (rewrite -> Nat.add_0_r). + rewrite <- fold_unfold_fib_nat_SS. + rewrite <- fold_unfold_fib_nat_SS. + reflexivity. +Qed. + +Compute (fib_pm22 nat Nat.add Nat.mul 0 1 0). (* = PM22 nat 1 0 0 1 *) +Compute (fib_pm22 nat Nat.add Nat.mul 0 1 1). (* = PM22 nat 0 1 1 1 *) +Compute (fib_pm22 nat Nat.add Nat.mul 0 1 2). (* = PM22 nat 1 1 1 2 *) +Compute (fib_pm22 nat Nat.add Nat.mul 0 1 3). (* = PM22 nat 1 2 2 3 *) +Compute (fib_pm22 nat Nat.add Nat.mul 0 1 4). (* = PM22 nat 2 3 3 5 *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (zero_pm22 nat 0) (one_pm22 nat 0 1) 0). +(* = PM22 (pm22 nat) (PM22 nat 1 0 0 1) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 1 0 0 1) *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (zero_pm22 nat 0) (one_pm22 nat 0 1) 1). +(* = PM22 (pm22 nat) (PM22 nat 0 0 0 0) (PM22 nat 1 0 0 1) (PM22 nat 1 0 0 1) (PM22 nat 1 0 0 1) *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (zero_pm22 nat 0) (one_pm22 nat 0 1) 2). +(* = PM22 (pm22 nat) (PM22 nat 1 0 0 1) (PM22 nat 1 0 0 1) (PM22 nat 1 0 0 1) (PM22 nat 2 0 0 2) *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (zero_pm22 nat 0) (one_pm22 nat 0 1) 3). +(* = PM22 (pm22 nat) (PM22 nat 1 0 0 1) (PM22 nat 2 0 0 2) (PM22 nat 2 0 0 2) (PM22 nat 3 0 0 3) *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (zero_pm22 nat 0) (one_pm22 nat 0 1) 4). +(* = PM22 (pm22 nat) (PM22 nat 2 0 0 2) (PM22 nat 3 0 0 3) (PM22 nat 3 0 0 3) (PM22 nat 5 0 0 5) *) + +Compute (fib_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul) (zero_pm22 nat 0) (one_pm22 nat 0 1) 5). +(* = PM22 (pm22 nat) (PM22 nat 3 0 0 3) (PM22 nat 5 0 0 5) (PM22 nat 5 0 0 5) (PM22 nat 8 0 0 8) *) + +Compute (fib_pm22 (pm22 (pm22 nat)) (add_pm22 (pm22 nat) (add_pm22 nat Nat.add)) (mul_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul)) (zero_pm22 (pm22 nat) (zero_pm22 nat 0)) (one_pm22 (pm22 nat) (zero_pm22 nat 0) (one_pm22 nat 0 1)) 0). +(* PM22 (pm22 (pm22 nat)) + (PM22 (pm22 nat) (PM22 nat 1 0 0 1) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 1 0 0 1)) + (PM22 (pm22 nat) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0)) + (PM22 (pm22 nat) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0)) + (PM22 (pm22 nat) (PM22 nat 1 0 0 1) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 1 0 0 1)) *) + +Compute (fib_pm22 (pm22 (pm22 nat)) (add_pm22 (pm22 nat) (add_pm22 nat Nat.add)) (mul_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul)) (zero_pm22 (pm22 nat) (zero_pm22 nat 0)) (one_pm22 (pm22 nat) (zero_pm22 nat 0) (one_pm22 nat 0 1)) 1). +(* PM22 (pm22 (pm22 nat)) + (PM22 (pm22 nat) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0)) + (PM22 (pm22 nat) (PM22 nat 1 0 0 1) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 1 0 0 1)) + (PM22 (pm22 nat) (PM22 nat 1 0 0 1) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 1 0 0 1)) + (PM22 (pm22 nat) (PM22 nat 1 0 0 1) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 1 0 0 1)) *) + +Compute (fib_pm22 (pm22 (pm22 nat)) (add_pm22 (pm22 nat) (add_pm22 nat Nat.add)) (mul_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul)) (zero_pm22 (pm22 nat) (zero_pm22 nat 0)) (one_pm22 (pm22 nat) (zero_pm22 nat 0) (one_pm22 nat 0 1)) 2). +(* PM22 (pm22 (pm22 nat)) + (PM22 (pm22 nat) (PM22 nat 1 0 0 1) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 1 0 0 1)) + (PM22 (pm22 nat) (PM22 nat 1 0 0 1) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 1 0 0 1)) + (PM22 (pm22 nat) (PM22 nat 1 0 0 1) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 1 0 0 1)) + (PM22 (pm22 nat) (PM22 nat 2 0 0 2) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 2 0 0 2)) *) + +Compute (fib_pm22 (pm22 (pm22 nat)) (add_pm22 (pm22 nat) (add_pm22 nat Nat.add)) (mul_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul)) (zero_pm22 (pm22 nat) (zero_pm22 nat 0)) (one_pm22 (pm22 nat) (zero_pm22 nat 0) (one_pm22 nat 0 1)) 3). +(* PM22 (pm22 (pm22 nat)) + (PM22 (pm22 nat) (PM22 nat 1 0 0 1) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 1 0 0 1)) + (PM22 (pm22 nat) (PM22 nat 2 0 0 2) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 2 0 0 2)) + (PM22 (pm22 nat) (PM22 nat 2 0 0 2) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 2 0 0 2)) + (PM22 (pm22 nat) (PM22 nat 3 0 0 3) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 3 0 0 3)) *) + +Compute (fib_pm22 (pm22 (pm22 nat)) (add_pm22 (pm22 nat) (add_pm22 nat Nat.add)) (mul_pm22 (pm22 nat) (add_pm22 nat Nat.add) (mul_pm22 nat Nat.add Nat.mul)) (zero_pm22 (pm22 nat) (zero_pm22 nat 0)) (one_pm22 (pm22 nat) (zero_pm22 nat 0) (one_pm22 nat 0 1)) 4). +(* = PM22 (pm22 (pm22 nat)) + (PM22 (pm22 nat) (PM22 nat 2 0 0 2) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 2 0 0 2)) + (PM22 (pm22 nat) (PM22 nat 3 0 0 3) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 3 0 0 3)) + (PM22 (pm22 nat) (PM22 nat 3 0 0 3) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 3 0 0 3)) + (PM22 (pm22 nat) (PM22 nat 5 0 0 5) (PM22 nat 0 0 0 0) (PM22 nat 0 0 0 0) (PM22 nat 5 0 0 5)) *) + +Proposition about_fib_pm22_ground : + forall n : nat, + fib_pm22 nat Nat.add Nat.mul 0 1 (S n) = + PM22 nat + (fib_nat n) (fib_nat (S n)) + (fib_nat (S n)) (fib_nat (S (S n))). +Proof. + intro n. + unfold fib_pm22. + induction n as [ | n' IHn']. + + - rewrite -> fold_unfold_power_pm22_S. + rewrite -> fold_unfold_power_pm22_O. + unfold F_pm22. + unfold one_pm22. + unfold mul_pm22. + repeat (rewrite -> Nat.add_0_l). + repeat (rewrite -> Nat.add_0_r). + repeat (rewrite -> Nat.mul_1_l). + rewrite -> fold_unfold_fib_nat_SS. + rewrite -> fold_unfold_fib_nat_S. + rewrite -> fold_unfold_fib_nat_O. + rewrite -> Nat.add_0_l. + reflexivity. + + - rewrite -> fold_unfold_power_pm22_S. + rewrite -> IHn'. + unfold mul_pm22, F_pm22. + rewrite ->2 Nat.mul_0_l. + rewrite ->2 Nat.add_0_l. + rewrite ->3 Nat.mul_1_l. + rewrite <-2 fold_unfold_fib_nat_SS. + reflexivity. +Qed. + +Proposition about_fib_pm22_base_case : + forall (V : Type) + (add mul : V -> V -> V) + (zero one : V), + (forall v : V, + add zero v = v) -> + (forall v : V, + add v zero = v) -> + (forall v : V, + mul zero v = zero) -> + (forall v : V, + mul v zero = zero) -> + (forall v : V, + mul one v = v) -> + (forall v : V, + mul v one = v) -> + forall (fib : nat -> V), + Fibonacci V add zero one fib -> + forall n : nat, + fib_pm22 V add mul (fib 0) (fib 1) (S n) = + PM22 V + (fib n) (fib (S n)) + (fib (S n)) (fib (S (S n))). +Proof. + intros V add mul zero one add_0_l add_0_r mul_0_l mul_0_r mul_1_l mul_1_r fib [H_fib_0 [H_fib_1 H_fib_SS]] n. + unfold fib_pm22. + induction n as [ | n' IHn']. + + - rewrite -> fold_unfold_power_pm22_S. + rewrite -> fold_unfold_power_pm22_O. + rewrite -> H_fib_SS. + rewrite -> H_fib_0. + rewrite -> H_fib_1. + rewrite -> add_0_l. + unfold F_pm22. + Check (mul_pm22_1_r V add mul zero add_0_l add_0_r mul_0_r one mul_1_r (PM22 V zero one one one)). + exact (mul_pm22_1_r V add mul zero add_0_l add_0_r mul_0_r one mul_1_r (PM22 V zero one one one)). + + - rewrite -> fold_unfold_power_pm22_S. + rewrite -> IHn'. + unfold F_pm22, mul_pm22. + rewrite -> H_fib_0. + rewrite -> H_fib_1. + rewrite ->2 mul_0_l. + rewrite ->3 mul_1_l. + rewrite ->2 add_0_l. + rewrite <-2 H_fib_SS. + reflexivity. +Qed. + +Proposition about_fib_pm22_ground' : + forall n : nat, + fib_pm22 nat Nat.add Nat.mul 0 1 (S n) = + PM22 nat + (fib_nat n) (fib_nat (S n)) + (fib_nat (S n)) (fib_nat (S (S n))). +Proof. + exact (about_fib_pm22_base_case nat Nat.add Nat.mul 0 1 Nat.add_0_l Nat.add_0_r Nat.mul_0_l Nat.mul_0_r Nat.mul_1_l Nat.mul_1_r fib_nat fib_nat_fits_Fibonacci). +Qed. + +(* ********** *) + +Definition scalar_mul_pm22 (V : Type) (mul : V -> V -> V) (v : V) (x : pm22 V) : pm22 V := + match x with + PM22 _ + x11 x12 + x21 x22 => + PM22 V + (mul v x11) (mul v x12) + (mul v x21) (mul v x22) + end. + +Proposition about_fib_pm22_induction_step : + forall (V : Type) + (add mul : V -> V -> V) + (zero one : V), + (forall v : V, + add zero v = v) -> + (forall v : V, + add v zero = v) -> + (forall v : V, + mul zero v = zero) -> + (forall v : V, + mul v zero = zero) -> + (forall v : V, + mul one v = v) -> + (forall v : V, + mul v one = v) -> + forall (fib : nat -> V), + Fibonacci V add zero one fib -> + forall n : nat, + fib_pm22 (pm22 V) (add_pm22 V add) (mul_pm22 V add mul) (zero_pm22 V (fib 0)) (one_pm22 V (fib 0) (fib 1)) (S n) = + PM22 (pm22 V) + (scalar_mul_pm22 V mul (fib n) (one_pm22 V zero one)) (scalar_mul_pm22 V mul (fib (S n)) (one_pm22 V zero one)) + (scalar_mul_pm22 V mul (fib (S n)) (one_pm22 V zero one)) (scalar_mul_pm22 V mul (fib (S (S n))) (one_pm22 V zero one)). +Proof. + intros V add mul zero one add_0_l add_0_r mul_0_l mul_0_r mul_1_l mul_1_r fib [H_fib_0 [H_fib_1 H_fib_SS]] n. + rewrite -> H_fib_0. + rewrite -> H_fib_1. + induction n as [ | n' IHn']. + + - Check (fold_unfold_fib_pm22_1 (pm22 V) (add_pm22 V add) (mul_pm22 V add mul) (zero_pm22 V zero) (add_pm22_0_l V add zero add_0_l) (add_pm22_0_r V add zero add_0_r) (mul_pm22_0_r V add mul zero add_0_r mul_0_r) (one_pm22 V zero one) (mul_pm22_1_r V add mul zero add_0_l add_0_r mul_0_r one mul_1_r)). + rewrite -> (fold_unfold_fib_pm22_1 (pm22 V) (add_pm22 V add) (mul_pm22 V add mul) (zero_pm22 V zero) (add_pm22_0_l V add zero add_0_l) (add_pm22_0_r V add zero add_0_r) (mul_pm22_0_r V add mul zero add_0_r mul_0_r) (one_pm22 V zero one) (mul_pm22_1_r V add mul zero add_0_l add_0_r mul_0_r one mul_1_r)). + unfold F_pm22, zero_pm22, one_pm22, scalar_mul_pm22. + rewrite ->3 mul_1_r. + rewrite ->3 mul_0_r. + repeat (rewrite -> H_fib_SS). + rewrite -> H_fib_1. + rewrite -> H_fib_0. + rewrite -> add_0_l. + reflexivity. + + - unfold fib_pm22. + unfold fib_pm22 in IHn'. + rewrite -> fold_unfold_power_pm22_S. + rewrite -> IHn'. + unfold F_pm22, zero_pm22, one_pm22, mul_pm22, add_pm22, scalar_mul_pm22. + repeat (rewrite -> add_0_l). + repeat (rewrite -> mul_0_l). + repeat (rewrite -> mul_1_l). + repeat (rewrite -> add_0_r). + repeat (rewrite -> add_0_l). + repeat (rewrite -> mul_1_r). + repeat (rewrite -> mul_0_r). + repeat (rewrite -> add_0_l). + rewrite <- H_fib_SS. + rewrite <- H_fib_SS. + reflexivity. +Qed. + +(* ********** *) + +(* end of week-10_a-fibonacci-structure.v *) diff --git a/cs3234/labs/week-10_the-abstraction-and-instantiation-of-Eureka-lemmas-about-resetting-the-accumulator.v b/cs3234/labs/week-10_the-abstraction-and-instantiation-of-Eureka-lemmas-about-resetting-the-accumulator.v new file mode 100644 index 0000000..bc5659d --- /dev/null +++ b/cs3234/labs/week-10_the-abstraction-and-instantiation-of-Eureka-lemmas-about-resetting-the-accumulator.v @@ -0,0 +1,434 @@ +(* week-10_the-abstraction-and-instantiation-of-Eureka-lemmas-about-resetting-the-accumulator.v *) +(* LPP 2024 - CS3234 2023-2024, Sem2 *) +(* Olivier Danvy *) +(* Version of 04 Apr 2024 *) + +(* ********** *) + +Ltac fold_unfold_tactic name := intros; unfold name; fold name; reflexivity. + +Require Import Arith Bool List. + +(* ********** *) + +Definition make_Eureka_lemma (A : Type) (id_A : A) (combine_A : A -> A -> A) (c : A -> A) (a : A) : Prop := + c a = combine_A (c id_A) a. + + +(* ********** *) + +Fixpoint power_acc (x n a : nat) : nat := + match n with + O => + a + | S n' => + power_acc x n' (x * a) + end. + +Definition power_alt (x n : nat) : nat := + power_acc x n 1. + + +Check (forall x n a : nat, make_Eureka_lemma nat 1 Nat.mul (power_acc x n) a). + +Lemma fold_unfold_power_acc_O : + forall x a : nat, + power_acc x 0 a = + a. +Proof. + fold_unfold_tactic power_acc. +Qed. + +Lemma fold_unfold_power_acc_S : + forall x n' a : nat, + power_acc x (S n') a = + power_acc x n' (x * a). +Proof. + fold_unfold_tactic power_acc. +Qed. + +Check (make_Eureka_lemma nat). +Check (make_Eureka_lemma nat 1). +Check (make_Eureka_lemma nat 1). +Check (make_Eureka_lemma nat 1 Nat.mul). +Check (forall x n a : nat, make_Eureka_lemma nat 1 Nat.mul (power_acc x n) a). + +Lemma about_power_acc : + forall x n a : nat, + power_acc x n a = power_acc x n 1 * a. +Proof. + intros x n. + induction n as [ | n' IHn']; intro a. + - rewrite ->2 fold_unfold_power_acc_O. + exact (eq_sym (Nat.mul_1_l a)). + - rewrite ->2 fold_unfold_power_acc_S. + rewrite -> Nat.mul_1_r. + rewrite -> (IHn' (x * a)). + rewrite -> (IHn' x). + Check Nat.mul_assoc. + apply Nat.mul_assoc. +Qed. + +Lemma about_power_acc_generically : + forall x n a : nat, + make_Eureka_lemma nat 1 Nat.mul (power_acc x n) a. +Proof. + unfold make_Eureka_lemma. + exact about_power_acc. +Qed. + +(* ********** *) + +Fixpoint add_acc (n a : nat) : nat := + match n with + O => + a + | S n' => + add_acc n' (S a) + end. + +Definition add_alt (n m : nat) : nat := + add_acc n m. + +Lemma fold_unfold_add_acc_O : + forall a : nat, + add_acc 0 a = + a. +Proof. + fold_unfold_tactic add_acc. +Qed. + +Lemma fold_unfold_add_acc_S : + forall n' a : nat, + add_acc (S n') a = + add_acc n' (S a). +Proof. + fold_unfold_tactic add_acc. +Qed. + +Lemma about_add_acc : + forall n a : nat, + add_acc n a = add_acc n 0 + a. +Proof. + intro n. + induction n as [ | n' IHn']; intro a. + - rewrite ->2 fold_unfold_add_acc_O. + exact (eq_sym (Nat.add_0_l a)). + - rewrite ->2 fold_unfold_add_acc_S. + rewrite -> (IHn' (S a)). + rewrite -> (IHn' 1). + rewrite <- Nat.add_assoc. + rewrite -> (Nat.add_1_l a). + reflexivity. +Qed. + +Lemma about_add_acc_generically : + forall n a : nat, + make_Eureka_lemma nat 0 Nat.add (add_acc n) a. +Proof. + unfold make_Eureka_lemma. + exact about_add_acc. +Qed. + +(* ********** *) + +(* Exercise 02 *) + +Fixpoint length_acc (V : Type) (vs : list V) (a : nat) : nat := + match vs with + nil => + a + | v :: vs' => + length_acc V vs' (S a) + end. + +Definition length_alt (V : Type) (vs : list V) : nat := + length_acc V vs 0. + +(* +Lemma about_length_acc : +*) + +(* +Lemma about_length_acc_generically : +*) + +(* ********** *) + +(* Exercise 03 *) + +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. + +Fixpoint reverse_acc (V : Type) (vs a : list V) : list V := + match vs with + nil => + a + | v :: vs' => + reverse_acc V vs' (v :: a) + end. + +Definition reverse_alt (V : Type) (vs : list V) : list V := + reverse_acc V vs nil. + +(* +Lemma about_reverse_acc : +*) + +(* +Lemma about_reverse_acc_generically : +*) + +(* ********** *) + +(* Exercise 04 *) + +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. + +(* +Lemma about_list_fold_left : +*) + +(* +Lemma about_list_fold_left_generically : +*) + +(* ********** *) + +(* Exercise 05 *) + +Fixpoint fac_acc (n a : nat) : nat := + match n with + | O => + a + | S n' => + fac_acc n' (S n' * a) + end. + +Definition fac_alt (n : nat) : nat := + fac_acc n 1. + +Lemma fold_unfold_fac_acc_O : + forall a : nat, + fac_acc 0 a = + a. +Proof. + fold_unfold_tactic fac_acc. +Qed. + +Lemma fold_unfold_fac_acc_S : + forall n' a : nat, + fac_acc (S n') a = + fac_acc n' (S n' * a). +Proof. + fold_unfold_tactic fac_acc. +Qed. + +(* +Lemma about_fac_acc : +*) + +(* +Lemma about_fac_acc_generically : +*) + +(* ********** *) + +(* Exercise 06 *) + +Inductive binary_tree : Type := +| Leaf : nat -> binary_tree +| Node : binary_tree -> binary_tree -> binary_tree. + +(* ***** *) + +Fixpoint weight (t : binary_tree) : nat := + match t with + | Leaf n => + n + | Node t1 t2 => + weight t1 + weight t2 + end. + +Lemma fold_unfold_weight_Leaf : + forall n : nat, + weight (Leaf n) = n. +Proof. + fold_unfold_tactic weight. +Qed. + +Lemma fold_unfold_weight_Node : + forall t1 t2 : binary_tree, + weight (Node t1 t2) = weight t1 + weight t2. +Proof. + fold_unfold_tactic weight. +Qed. + +(* ***** *) + +Fixpoint weight_acc (t : binary_tree) (a : nat) : nat := + match t with + | Leaf n => + n + a + | Node t1 t2 => + weight_acc t1 (weight_acc t2 a) + end. + +Definition weight_alt (t : binary_tree) : nat := + weight_acc t 0. + +Lemma fold_unfold_weight_acc_Leaf : + forall n a : nat, + weight_acc (Leaf n) a = n + a. +Proof. + fold_unfold_tactic weight_acc. +Qed. + +Lemma fold_unfold_weight_acc_Node : + forall (t1 t2 : binary_tree) + (a : nat), + weight_acc (Node t1 t2) a = weight_acc t1 (weight_acc t2 a). +Proof. + fold_unfold_tactic weight_acc. +Qed. + +(* Nat fold right is a generic version of the non accumulator based iteration function *) +Definition nat_fold_right (V: Type) (zero_case: V) (succ_case: V -> V) (n : nat) : V := + let fix visit i := + match i with + O => + zero_case + | S n' => + succ_case (visit n') + end + in visit n. + +(* Nat fold left is a generic version of the accumulator based iteration function *) +(* This is tail recursive ? *) +Definition nat_fold_left (V: Type) (zero_case: V) (succ_case: V -> V) (n : nat) : V := + let fix visit i a := + match i with + O => + a + | S n' => + visit n' (succ_case a) + end + in visit n zero_case. + +Definition nat_fold_left_with_right_1 (V: Type) (zero_case: V) (succ_case: V -> V) (n : nat) : V := + let fix visit i := fun a => + match i with + O => + a + | S n' => + visit n' (succ_case a) + end + in visit n zero_case. + +Definition nat_fold_left_with_right_2 (V: Type) (zero_case: V) (succ_case: V -> V) (n : nat) : V := + let fix visit i := + match i with + O => + fun a => a + | S n' => + fun a => visit n' (succ_case a) + end + in visit n zero_case. + +Definition nat_fold_left_with_right_3 (V: Type) (zero_case: V) (succ_case: V -> V) (n : nat) : V := + (nat_fold_right + (V -> V) + (fun a => a) + (fun ih => fun a => ih (succ_case a)) + n) + zero_case. + +Definition test_add (candidate : nat -> nat -> nat) := + (Nat.eqb (candidate 0 0) 0) && + (Nat.eqb (candidate 1 0) 1) && + (Nat.eqb (candidate 0 1) 1) && + (Nat.eqb (candidate 1 1) 2) && + (Nat.eqb (candidate 1 2) 3) && + (Nat.eqb (candidate 2 2) 4). + +Definition nat_add (a b : nat) := + let fix visit i := + match i with + | O => b + | S n' => S (visit n') + end + in visit a. + +Definition nat_add_acc (m n : nat) := + let fix visit i a := + match i with + | O => m + | S n' => visit n' (S a) + end + in visit m n. + + +Compute test_add Nat.add. +Compute test_add nat_add. +Compute test_add nat_add_acc. + +Compute test_add (fun a b => nat_fold_left_with_right_3 + nat + b + S + a). + +Definition nat_parafold_right (V: Type) (zero_case: V) (succ_case: nat -> V -> V) (n : nat) : V := + let fix visit i := + match i with + O => + zero_case + | S i' => + succ_case i' (visit i') + end + in visit n. + +Definition nat_parafold_right_1 (V: Type) (zero_case: V) (succ_case: nat -> V -> V) (n : nat) : V := + nat_fold_right + + +Definition r_fac (n : nat) : nat := + let fix visit i := + match i with + | O => 1 + | S i' => match i' with + | O => 1 + | S i'' => visit + +(* Nat fold left is a generic version of the accumulator based iteration function *) +Definition nat_fold_left (V: Type) (zero_case: V) (succ_case: V -> V) (n : nat) : V := + let fix visit i a := + match i with + O => + a + | S n' => + visit n' (succ_case a) + end + in visit n zero_case. +(* +Lemma about_weight_acc : +*) + +(* +Lemma about_weight_acc_generically : +*) + +(* ********** *) + +(* end of week-10_the-abstraction-and-instantiation-of-Eureka-lemmas-about-resetting-the-accumulator.v *) diff --git a/cs3234/labs/week-11_induction-principles.v b/cs3234/labs/week-11_induction-principles.v new file mode 100644 index 0000000..12e6c99 --- /dev/null +++ b/cs3234/labs/week-11_induction-principles.v @@ -0,0 +1,494 @@ +(* week-11_induction-principles.v *) +(* LPP 2024 - CS3236 2023-2024, Sem2 *) +(* Olivier Danvy *) +(* Version of 05 Apr 2024 *) + +(* ********** *) + +Ltac fold_unfold_tactic name := intros; unfold name; fold name; reflexivity. + +Require Import Arith Bool. + +(* ********** *) + +(* One goal of this lecture is to revisit the proof that + at most one function satisfies the following specification. +*) + +Definition specification_of_the_fibonacci_function (fib : nat -> nat) := + fib 0 = 0 + /\ + fib 1 = 1 + /\ + forall n'' : nat, + fib (S (S n'')) = fib n'' + fib (S n''). + +Fixpoint fib (n : nat) : nat := + match n with + | 0 => + 0 + | S n' => + match n' with + | 0 => + 1 + | S n'' => + fib n'' + fib n' + end + end. + +Lemma fold_unfold_fib_O : + fib 0 = + 0. +Proof. + fold_unfold_tactic fib. +Qed. + +Lemma fold_unfold_fib_S : + forall n' : nat, + fib (S n') = + match n' with + | 0 => + 1 + | S n'' => + fib n'' + fib n' + end. +Proof. + fold_unfold_tactic fib. +Qed. + +Corollary fold_unfold_fib_1 : + fib 1 = + 1. +Proof. + rewrite -> fold_unfold_fib_S. + reflexivity. +Qed. + +Corollary fold_unfold_fib_SS : + forall n'' : nat, + fib (S (S n'')) = + fib n'' + fib (S n''). +Proof. + intro n''. + rewrite -> fold_unfold_fib_S. + reflexivity. +Qed. + +Proposition fib_satisfies_the_specification_of_fib : + specification_of_the_fibonacci_function fib. +Proof. + unfold specification_of_the_fibonacci_function. + Check (conj fold_unfold_fib_O (conj fold_unfold_fib_1 fold_unfold_fib_SS)). + exact (conj fold_unfold_fib_O (conj fold_unfold_fib_1 fold_unfold_fib_SS)). +Qed. + +(* ********** *) + +(* The mathematical induction principle already exists, + it is the structural induction principle associated to Peano numbers: +*) + +Check nat_ind. + +(* But we can still express it ourselves. + We can also prove it using the resident mathematical induction principle, + either implicitly or explicitly: +*) + +Lemma nat_ind1 : + forall P : nat -> Prop, + P 0 -> + (forall n : nat, P n -> P (S n)) -> + forall n : nat, P n. +Proof. +Abort. + +(* ********** *) + +(* We can also use nat_ind as an ordinary lemma + instead of using the induction tactic: +*) + +Fixpoint r_add (i j : nat) : nat := + match i with + | O => + j + | S i' => + S (r_add i' j) + end. + +Lemma fold_unfold_r_add_O : + forall j : nat, + r_add 0 j = + j. +Proof. + fold_unfold_tactic r_add. +Qed. + +Lemma fold_unfold_r_add_S : + forall i' j : nat, + r_add (S i') j = + S (r_add i' j). +Proof. + fold_unfold_tactic r_add. +Qed. + +Proposition r_add_0_r : + forall i : nat, + r_add i 0 = i. +Proof. + (* First, a routine induction: *) + intro i. + induction i as [ | i' IHi']. + - exact (fold_unfold_r_add_O 0). + - rewrite -> fold_unfold_r_add_S. + Check f_equal. + Check (f_equal S). (* : forall x y : nat, x = y -> S x = S y *) + Check (f_equal S IHi'). + exact (f_equal S IHi'). + + Restart. + + (* And now for using nat_ind: *) + Check nat_ind. +Abort. + +(* ********** *) + +Fixpoint fibfib (n : nat) : nat * nat := + match n with + | O => + (0, 1) + | S n' => + let (fib_n', fib_succ_n') := fibfib n' + in (fib_succ_n', fib_n' + fib_succ_n') + end. + +Definition fib_lin (n : nat) : nat := + let (fib_n, _) := fibfib n + in fib_n. + +Lemma fold_unfold_fibfib_O : + fibfib 0 = + (0, 1). +Proof. + fold_unfold_tactic fibfib. +Qed. + +Lemma fold_unfold_fibfib_S : + forall n' : nat, + fibfib (S n') = + let (fib_n', fib_succ_n') := fibfib n' + in (fib_succ_n', fib_n' + fib_succ_n'). +Proof. + fold_unfold_tactic fibfib. +Qed. + +Lemma about_fibfib : + forall fib : nat -> nat, + specification_of_the_fibonacci_function fib -> + forall n : nat, + fibfib n = (fib n, fib (S n)). +Proof. + unfold specification_of_the_fibonacci_function. + intros fib [S_fib_O [S_fib_1 S_fib_SS]] n. + induction n as [ | [ | n''] IH]. + - rewrite -> fold_unfold_fibfib_O. + rewrite -> S_fib_O. + rewrite -> S_fib_1. + reflexivity. + - rewrite -> fold_unfold_fibfib_S. + rewrite -> fold_unfold_fibfib_O. + rewrite -> S_fib_1. + rewrite -> S_fib_SS. + rewrite -> S_fib_O. + rewrite -> S_fib_1. + reflexivity. + - rewrite -> fold_unfold_fibfib_S. + rewrite -> IH. + rewrite <- (S_fib_SS (S n'')). + reflexivity. +Qed. + +Proposition fib_lin_satisfies_the_specification_of_fib : + specification_of_the_fibonacci_function fib_lin. +Proof. + unfold specification_of_the_fibonacci_function, fib_lin. + split. + - rewrite -> fold_unfold_fibfib_O. + reflexivity. + - split. + + rewrite -> fold_unfold_fibfib_S. + rewrite -> fold_unfold_fibfib_O. + reflexivity. + + intro i. + Check (about_fibfib fib fib_satisfies_the_specification_of_fib (S (S i))). + rewrite -> (about_fibfib fib fib_satisfies_the_specification_of_fib (S (S i))). + rewrite -> (about_fibfib fib fib_satisfies_the_specification_of_fib i). + rewrite -> (about_fibfib fib fib_satisfies_the_specification_of_fib (S i)). + exact (fold_unfold_fib_SS i). +Qed. + +(* ********** *) + +(* We can also express a mathematical induction principle + with two base cases and two induction hypotheses + that befits the structure of the Fibonacci function: +*) + +Lemma nat_ind2 : + forall P : nat -> Prop, + P 0 -> + P 1 -> + (forall n : nat, P n -> P (S n) -> P (S (S n))) -> + forall n : nat, P n. +Proof. + intros P H_P0 H_P1 H_PSS n. + induction n as [ | [ | n''] IHn']. +Abort. + +(* Thus equipped, the following theorem is proved pretty directly: *) + +Theorem there_is_at_most_one_fibonacci_function : + forall fib1 fib2 : nat -> nat, + specification_of_the_fibonacci_function fib1 -> + specification_of_the_fibonacci_function fib2 -> + forall n : nat, + fib1 n = fib2 n. +Proof. + intros fib1 fib2. + unfold specification_of_the_fibonacci_function. + intros [H_fib1_0 [H_fib1_1 H_fib1_SS]] + [H_fib2_0 [H_fib2_1 H_fib2_SS]] + n. + induction n as [ | | n'' IHn'' IHSn''] using nat_ind2. +Abort. + +(* ***** *) + +Fixpoint evenp1 (n : nat) : bool := + match n with + | 0 => + true + | S n' => + negb (evenp1 n') + end. + +Lemma fold_unfold_evenp1_O : + evenp1 0 = + true. +Proof. + fold_unfold_tactic evenp1. +Qed. + +Lemma fold_unfold_evenp1_S : + forall n' : nat, + evenp1 (S n') = + negb (evenp1 n'). +Proof. + fold_unfold_tactic evenp1. +Qed. + +(* ***** *) + +(* The evenness predicate is often programmed tail-recursively + and with no accumulator, by peeling two layers of S at a time. + Its equivalence with evenp1 is messy to prove by mathematical induction + but effortless using nat_ind2: +*) + +Fixpoint evenp2 (n : nat) : bool := + match n with + | 0 => + true + | S n' => + match n' with + | 0 => + false + | S n'' => + evenp2 n'' + end + end. + +Lemma fold_unfold_evenp2_O : + evenp2 0 = + true. +Proof. + fold_unfold_tactic evenp2. +Qed. + +Lemma fold_unfold_evenp2_S : + forall n' : nat, + evenp2 (S n') = + match n' with + | 0 => + false + | S n'' => + evenp2 n'' + end. +Proof. + fold_unfold_tactic evenp2. +Qed. + +Corollary fold_unfold_evenp2_1 : + evenp2 1 = + false. +Proof. + rewrite -> fold_unfold_evenp2_S. + reflexivity. +Qed. + +Corollary fold_unfold_evenp2_SS : + forall n'' : nat, + evenp2 (S (S n'')) = + evenp2 n''. +Proof. + intro n''. + rewrite -> fold_unfold_evenp2_S. + reflexivity. +Qed. + +Theorem evenp1_and_evenp2_are_functionally_equal : + forall n : nat, + evenp1 n = evenp2 n. +Proof. + intro n. + induction n as [ | n' IHn']. +Abort. + +(* ***** *) + +Lemma two_times_S : + forall n : nat, + S (S (2 * n)) = 2 * S n. +Proof. +Admitted. + +Theorem soundness_and_completeness_of_evenp_using_nat_ind2 : + forall n : nat, + evenp2 n = true <-> exists m : nat, n = 2 * m. +Proof. + intro n. + induction n as [ | | n' [IHn'_sound IHn'_complete] [IHSn'_sound IHSn'_complete]] using nat_ind2. +Abort. + +(* ***** *) + +(* For another example, we can prove the mathematical induction principle using nat_ind2: *) + +Lemma nat_ind1' : + forall P : nat -> Prop, + P 0 -> + (forall n : nat, P n -> P (S n)) -> + forall n : nat, P n. +Proof. + intros P H_P0 H_PS n. + induction n as [ | | n' IHn'] using nat_ind2. + - exact H_P0. + - Check (H_PS 0 H_P0). + exact (H_PS 0 H_P0). + - Check (H_PS (S n') IHn). + exact (H_PS (S n') IHn). +Qed. + +(* We can also generalize nat_ind2 to an induction principle + with three base cases and three induction hypotheses: *) + +Lemma nat_ind3 : + forall P : nat -> Prop, + P 0 -> + P 1 -> + P 2 -> + (forall n : nat, P n -> P (S n) -> P (S (S n)) -> P (S (S (S n)))) -> + forall n : nat, P n. +Proof. +Abort. + +(* ***** *) + +Fixpoint ternaryp (n : nat) : bool := + match n with + | 0 => + true + | 1 => + false + | 2 => + false + | S (S (S n')) => + ternaryp n' + end. + +Lemma fold_unfold_ternaryp_O : + ternaryp 0 = + true. +Proof. + fold_unfold_tactic ternaryp. +Qed. + +Lemma fold_unfold_ternaryp_1 : + ternaryp 1 = + false. +Proof. + fold_unfold_tactic ternaryp. +Qed. + +Lemma fold_unfold_ternaryp_2 : + ternaryp 2 = + false. +Proof. + fold_unfold_tactic ternaryp. +Qed. + +Lemma fold_unfold_ternaryp_SSS : + forall n' : nat, + ternaryp (S (S (S n'))) = + ternaryp n'. +Proof. + fold_unfold_tactic ternaryp. +Qed. + +Theorem soundness_and_completeness_of_ternaryp_using_nat_ind3 : + forall n : nat, + ternaryp n = true <-> exists m : nat, n = 3 * m. +Proof. +Abort. + +(* ********** *) + +Lemma three_times_S : + forall n : nat, + S (S (S (3 * n))) = 3 * S n. +Proof. +Admitted. + +Property threes_and_fives : + forall n : nat, + exists a b : nat, + 8 + n = 3 * a + 5 * b. +Proof. +Abort. + +(* ********** *) + +Lemma four_times_S : + forall n : nat, + S (S (S (S (4 * n)))) = 4 * S n. +Proof. +Admitted. + +Lemma five_times_S : + forall n : nat, + S (S (S (S (S (5 * n))))) = 5 * S n. +Proof. +Admitted. + +Property fours_and_fives : + forall n : nat, + exists a b : nat, + 12 + n = 4 * a + 5 * b. +Proof. +Abort. + +(* ********** *) + +(* end of week-10_induction-principles.v *) diff --git a/cs3234/labs/week-13_a-continuation-based-interpreter.v b/cs3234/labs/week-13_a-continuation-based-interpreter.v new file mode 100644 index 0000000..d998be5 --- /dev/null +++ b/cs3234/labs/week-13_a-continuation-based-interpreter.v @@ -0,0 +1,193 @@ +(* week-13_a-continuation-based-interpreter.v *) +(* LPP 2024 - CS3234 2023-2024, Sem2 *) +(* Olivier Danvy *) +(* Version of 12 Apr 2024 *) + +(* ********** *) + +Ltac fold_unfold_tactic name := intros; unfold name; fold name; reflexivity. + +Require Import Arith Bool List String Ascii. + +(* ***** *) + +Check (1 =? 2). +Check Nat.eqb. +Check (Nat.eqb 1 2). + +Check (1 <=? 2). +Check Nat.leb. +Check (Nat.leb 1 2). + +Check (1 arithmetic_expression +| Plus : arithmetic_expression -> arithmetic_expression -> arithmetic_expression +| Minus : arithmetic_expression -> arithmetic_expression -> arithmetic_expression. + +(* Source programs: *) + +Inductive source_program : Type := + Source_program : arithmetic_expression -> source_program. + +(* ********** *) + +(* Semantics: *) + +Inductive expressible_value : Type := + Expressible_nat : nat -> expressible_value +| Expressible_msg : string -> expressible_value. + +(* ********** *) + +Definition specification_of_evaluate (evaluate : arithmetic_expression -> expressible_value) := + (forall n : nat, + evaluate (Literal n) = Expressible_nat n) + /\ + ((forall (ae1 ae2 : arithmetic_expression) + (s1 : string), + evaluate ae1 = Expressible_msg s1 -> + evaluate (Plus ae1 ae2) = Expressible_msg s1) + /\ + (forall (ae1 ae2 : arithmetic_expression) + (n1 : nat) + (s2 : string), + evaluate ae1 = Expressible_nat n1 -> + evaluate ae2 = Expressible_msg s2 -> + evaluate (Plus ae1 ae2) = Expressible_msg s2) + /\ + (forall (ae1 ae2 : arithmetic_expression) + (n1 n2 : nat), + evaluate ae1 = Expressible_nat n1 -> + evaluate ae2 = Expressible_nat n2 -> + evaluate (Plus ae1 ae2) = Expressible_nat (n1 + n2))) + /\ + ((forall (ae1 ae2 : arithmetic_expression) + (s1 : string), + evaluate ae1 = Expressible_msg s1 -> + evaluate (Minus ae1 ae2) = Expressible_msg s1) + /\ + (forall (ae1 ae2 : arithmetic_expression) + (n1 : nat) + (s2 : string), + evaluate ae1 = Expressible_nat n1 -> + evaluate ae2 = Expressible_msg s2 -> + evaluate (Minus ae1 ae2) = Expressible_msg s2) + /\ + (forall (ae1 ae2 : arithmetic_expression) + (n1 n2 : nat), + evaluate ae1 = Expressible_nat n1 -> + evaluate ae2 = Expressible_nat n2 -> + n1 + evaluate (Minus ae1 ae2) = Expressible_msg (String.append "numerical underflow: -" (string_of_nat (n2 - n1)))) + /\ + (forall (ae1 ae2 : arithmetic_expression) + (n1 n2 : nat), + evaluate ae1 = Expressible_nat n1 -> + evaluate ae2 = Expressible_nat n2 -> + n1 + evaluate (Minus ae1 ae2) = Expressible_nat (n1 - n2))). + +Definition specification_of_interpret (interpret : source_program -> expressible_value) := + forall evaluate : arithmetic_expression -> expressible_value, + specification_of_evaluate evaluate -> + forall ae : arithmetic_expression, + interpret (Source_program ae) = evaluate ae. + +(* ********** *) + +(* An evaluation function in delimited continuation-passing style: *) + +Fixpoint evaluate_cb (ae : arithmetic_expression) (k : nat -> expressible_value) : expressible_value := + Expressible_msg "not implemented yet". + +(* +Lemma fold_unfold_evaluate_cb_Literal : + forall (n : nat) + (k : nat -> expressible_value), + evaluate_cb (Literal n) k = + ... +Proof. + fold_unfold_tactic evaluate_cb. +Qed. + +Lemma fold_unfold_evaluate_cb_Plus : + forall (ae1 ae2 : arithmetic_expression) + (k : nat -> expressible_value), + evaluate_cb (Plus ae1 ae2) k = + ... +Proof. + fold_unfold_tactic evaluate_cb. +Qed. + +Lemma fold_unfold_evaluate_cb_Minus : + forall (ae1 ae2 : arithmetic_expression) + (k : nat -> expressible_value), + evaluate_cb (Minus ae1 ae2) k = + ... +Proof. + fold_unfold_tactic evaluate_cb. +Qed. +*) + +(* ***** *) + +Definition interpret (sp : source_program) : expressible_value := + Expressible_msg "not implemented yet". + +(* ***** *) + +Lemma about_evaluate_cb : + forall evaluate : arithmetic_expression -> expressible_value, + specification_of_evaluate evaluate -> + forall ae : arithmetic_expression, + (exists n : nat, + evaluate ae = Expressible_nat n + /\ + forall k : nat -> expressible_value, + evaluate_cb ae k = k n) + \/ + (exists s : string, + evaluate ae = Expressible_msg s + /\ + forall k : nat -> expressible_value, + evaluate_cb ae k = Expressible_msg s). +Proof. +Abort. + +(* ***** *) + +Theorem interpret_satisfies_specification_of_interpret : + specification_of_interpret interpret. +Proof. +Abort. + +(* ********** *) + +(* end of week-13_a-continuation-based-interpreter.v *)