nus/cs3234/labs/week-02_formalizing-a-proof-in-tcpa.v
2024-02-04 15:40:50 +08:00

326 lines
6.7 KiB
Coq

(* week-02_formalizing-a-proof-in-tcpa.v *)
(* LPP 2024 - CS3234 2023-2024, Sem2 *)
(* Olivier Danvy <danvy@yale-nus.edu.sg> *)
(* Version of 26 Jan 2024 with a correct definition of the specifications *)
(* was: *)
(* Version of 25 Jan 2024 *)
(* ********** *)
Require Import Arith Bool.
(* ********** *)
Definition lambda_dropped_specification_of_the_addition_function (add : nat -> nat -> nat) : Prop :=
forall j : nat,
(add O j = j)
/\
(forall i' ih : nat,
add i' j = ih ->
add (S i') j = S ih).
Definition lambda_lifted_specification_of_the_addition_function (add : nat -> nat -> nat) : Prop :=
(forall j : nat,
add O j = j)
/\
(forall j i' ih : nat,
add i' j = ih ->
add (S i') j = S ih).
(* ********** *)
Lemma soundness_of_evenp (evenp : nat -> bool) :
forall n : nat,
evenp n = true ->
exists m : nat,
n = 2 * m.
Admitted.
Lemma completeness_of_evenp (evenp : nat -> bool) :
forall n : nat,
(exists m : nat,
n = 2 * m) ->
evenp n = true.
Admitted.
(* ********** *)
(* Exercise 01 *)
Proposition soundness_of_the_lambda_dropped_specification_of_the_addition_function :
forall add : nat -> nat -> nat,
lambda_dropped_specification_of_the_addition_function add ->
forall x y z : nat,
add x y = z ->
z = x + y.
Admitted.
Proposition completeness_of_the_lambda_lifted_specification_of_the_addition_function :
forall add : nat -> nat -> nat,
lambda_lifted_specification_of_the_addition_function add ->
forall x y z : nat,
x + y = z ->
z = add x y.
Admitted.
(* ********** *)
(* Exercise 02 *)
Definition specification_of_negation (candidate : bool -> bool) : Prop :=
(candidate true = false) /\
(candidate false = true).
Theorem soundness_of_the_specification_of_negation_with_respect_to_negb :
forall candidate: bool -> bool,
specification_of_negation candidate ->
forall b b': bool,
candidate b = b' ->
b' = negb b.
Admitted.
Theorem completeness_of_the_specification_of_negation_with_respect_to_negb :
forall candidate: bool -> bool,
specification_of_negation candidate ->
forall b b': bool,
negb b = b' ->
candidate b = b'.
Admitted.
Definition negation (b : bool) : bool :=
if b then false else true.
Theorem negation_satisfies_the_specification_of_negation :
specification_of_negation negation.
Admitted.
Theorem soundness_of_negation_with_respect_to_negb :
forall b b' : bool,
negation b = b' ->
b' = negb b.
Admitted.
Theorem completeness_of_negation_with_respect_to_negb :
forall b b' : bool,
negb b = b' ->
b' = negation b.
Admitted.
Theorem soundness_of_negb_with_respect_to_negation :
forall b b' : bool,
negb b = b' ->
b' = negation b.
Admitted.
Theorem completeness_of_negb_with_respect_to_negation :
forall b b' : bool,
negation b = b' ->
b' = negb b.
Admitted.
(* ********** *)
Search (_ + 0 = _).
Check Nat.add_0_r.
Proposition first_formal_proof :
forall n : nat,
n + 0 = 0 + n.
Proof.
intro n.
Check (Nat.add_0_r n).
rewrite -> (Nat.add_0_r n).
Check (Nat.add_0_l n).
rewrite -> (Nat.add_0_l n).
reflexivity.
Qed.
Check first_formal_proof.
Proposition first_formal_proof' :
forall n : nat,
n + 0 = 0 + n.
Proof.
intro n.
rewrite -> (Nat.add_0_l n).
rewrite -> (Nat.add_0_r n).
reflexivity.
Qed.
Proposition first_formal_proof'' :
forall n : nat,
n + 0 = 0 + n.
Proof.
intro n.
rewrite -> (Nat.add_0_r n).
rewrite -> (Nat.add_0_l n).
reflexivity.
Restart.
intro n.
rewrite -> (Nat.add_0_l n).
rewrite -> (Nat.add_0_r n).
reflexivity.
Qed.
Proposition first_formal_proof''' :
forall n : nat,
n + 0 = 0 + n.
Proof.
intro n.
Check (Nat.add_comm n 0).
exact (Nat.add_comm n 0).
Restart.
intro n.
Check (Nat.add_comm 0 n).
symmetry.
exact (Nat.add_comm 0 n).
Qed.
(* ********** *)
Search (_ + _ = _ + _).
Proposition first_formal_proof'''' :
forall n : nat,
n + 0 = 0 + n.
Proof.
intro n.
exact (Nat.add_comm n 0).
Restart.
intro n.
symmetry.
exact (Nat.add_comm 0 n).
Qed.
(*
Nat.add_comm: forall n m : nat, n + m = m + n
Nat.add_assoc: forall n m p : nat, n + (m + p) = n + m + p
*)
Check (Nat.add_comm 2 3).
Proposition second_formal_proof:
forall x y z : nat,
x + (y + z) = y + (x + z).
Proof.
intro x.
intro y.
intro z.
Search (_ + (_ + _)).
Check (Nat.add_assoc x y z).
rewrite -> (Nat.add_assoc x y z).
Check Nat.add_comm.
rewrite -> (Nat.add_comm x y).
rewrite <- (Nat.add_assoc).
reflexivity.
Restart.
intros x y z.
rewrite -> (Nat.add_assoc x y z).
rewrite -> (Nat.add_comm x y).
rewrite <- (Nat.add_assoc y x z).
reflexivity.
Qed.
(* ********** *)
Proposition third_formal_proof:
forall n : nat,
n + 0 + 0 = 0 + 0 + n.
Proof.
intro n.
rewrite -> (Nat.add_0_r n).
rewrite -> (Nat.add_0_r n).
Check (Nat.add_0_l 0).
rewrite -> (Nat.add_0_l 0).
rewrite -> (Nat.add_0_l n).
reflexivity.
Restart.
intro n.
rewrite -> (Nat.add_0_r n).
rewrite -> (Nat.add_0_r n).
rewrite -> (Nat.add_0_l 0).
exact (Nat.add_0_l n).
Restart.
intro n.
Check (first_formal_proof (n + 0)).
rewrite -> (first_formal_proof (n + 0)).
Check (first_formal_proof n).
rewrite -> (first_formal_proof n).
exact (Nat.add_assoc 0 0 n).
Restart.
intro n.
Check (Nat.add_0_l 0).
rewrite -> (Nat.add_0_l 0).
rewrite -> (Nat.add_0_r n).
Check first_formal_proof.
exact (first_formal_proof n).
Qed.
(* ********** *)
(* Exercise 3 *)
Proposition foo :
forall n1 n2 n3 n4 : nat,
n1 + (n2 + (n3 + n4)) = ((n1 + n2) + n3) + n4.
Proof.
intros a b c d.
Search (_ + (_ + _) = _ + _ + _).
rewrite -> (Nat.add_assoc a b (c + d)).
Check Nat.add_assoc.
rewrite -> (Nat.add_assoc (a + b) c d).
reflexivity.
Qed.
(* ********** *)
(* Exercise 4 *)
(* ********** *)
(* Exercise 5 *)
Proposition binomial_expansion_2_warmup :
forall x y : nat,
(x + y) * (x + y) = x * x + x * y + x * y + y * y.
Proof.
intros x y.
Search (_ * (_ + _) = _).
Check Nat.mul_add_distr_l.
Check Nat.mul_add_distr_r.
rewrite -> (Nat.mul_add_distr_l (x + y) x y).
rewrite -> (Nat.mul_add_distr_r x y x ).
rewrite -> (Nat.mul_comm y x).
rewrite -> (Nat.mul_add_distr_r x y y ).
Check Nat.add_assoc.
rewrite -> (Nat.add_assoc (x * x + x * y) (x * y) (y * y)).
reflexivity.
Qed.
Proposition binomial_expansion_2 :
forall x y : nat,
(x + y) * (x + y) = x * x + 2 * (x * y) + y * y.
Proof.
intros x y.
Check (Nat.mul_succ_l 1 (x * y)).
rewrite -> (Nat.mul_succ_l 1 (x * y)).
Search (1 * _).
rewrite -> (Nat.mul_1_l (x * y)).
rewrite -> (Nat.add_assoc (x * x) (x * y) (x * y)).
exact (binomial_expansion_2_warmup x y).
Qed.
(* ********** *)
(* end of week-02_formalizing-a-proof-in-tcpa.v *)