nus/cs3234/labs/week-03_structural-induction-over-Peano-numbers.v
2024-02-13 17:04:53 +08:00

135 lines
2.9 KiB
Coq

(* week-03_structural-induction-over-Peano-numbers.v *)
(* LPP 2023 - CS3234 2023-2024, Sem2 *)
(* Olivier Danvy <danvy@yale-nus.edu.sg> *)
(* Version of 05 Feb 2024, with a revisitation of Proposition add_0_r_v0 that uses the induction tactic *)
(* was: *)
(* Version of 02 Feb 2024 *)
(* ********** *)
Definition specification_of_the_addition_function_1 (add : nat -> nat -> nat) :=
(forall j : nat,
add O j = j)
/\
(forall i' j : nat,
add (S i') j = S (add i' j)).
Check nat_ind.
(*
nat_ind
: forall P : nat -> Prop, P 0 -> (forall n : nat, P n -> P (S n)) -> forall n : nat, P n
*)
Proposition add_0_r_v0 :
forall add : nat -> nat -> nat,
specification_of_the_addition_function_1 add ->
forall n : nat,
add n 0 = n.
Proof.
intro add.
unfold specification_of_the_addition_function_1.
intros [S_add_O S_add_S].
Check nat_ind.
Check (nat_ind
(fun n => add n 0 = n)).
Check (S_add_O 0).
Check (nat_ind
(fun n => add n 0 = n)
(S_add_O 0)).
apply (nat_ind
(fun n => add n 0 = n)
(S_add_O 0)).
intro n'.
intro IHn'.
Check (S_add_S n' 0).
rewrite -> (S_add_S n' 0).
rewrite -> IHn'.
reflexivity.
Qed.
Search (0 + _ = _).
(* plus_O_n: forall n : nat, 0 + n = n *)
Search (S _ + _ = _).
(* plus_Sn_m: forall n m : nat, S n + m = S (n + m) *)
Proposition add_0_r_v1 :
forall n : nat,
n + 0 = n.
Proof.
Check nat_ind.
Check (nat_ind
(fun n => n + 0 = n)).
Check plus_O_n.
Check (plus_O_n 0).
Check (nat_ind
(fun n => n + 0 = n)
(plus_O_n 0)).
apply (nat_ind
(fun n => n + 0 = n)
(plus_O_n 0)).
intros n'.
intros IHn'.
Check (plus_Sn_m n' 0).
rewrite -> (plus_Sn_m n' 0).
rewrite -> IHn'.
reflexivity.
Restart.
intro n.
induction n using nat_ind.
- Check (plus_O_n 0).
exact (plus_O_n 0).
- revert n IHn.
intros n' IHn'.
Check (plus_Sn_m n' 0).
rewrite -> (plus_Sn_m n' 0).
rewrite -> IHn'.
reflexivity.
Restart.
intro n.
induction n as [ | n' IHn'] using nat_ind.
- Check (plus_O_n 0).
exact (plus_O_n 0).
- Check (plus_Sn_m n' 0).
rewrite -> (plus_Sn_m n' 0).
rewrite -> IHn'.
reflexivity.
Restart.
intro n.
induction n as [ | n' IHn'].
- Check (plus_O_n 0).
exact (plus_O_n 0).
- Check (plus_Sn_m n' 0).
rewrite -> (plus_Sn_m n' 0).
rewrite -> IHn'.
reflexivity.
Qed.
Proposition add_0_r_v0_revisited :
forall add : nat -> nat -> nat,
specification_of_the_addition_function_1 add ->
forall n : nat,
add n 0 = n.
Proof.
intro add.
unfold specification_of_the_addition_function_1.
intros [S_add_O S_add_S].
intro n.
induction n as [ | n' IHn'].
- exact (S_add_O 0).
- rewrite -> (S_add_S n' 0).
rewrite -> IHn'.
reflexivity.
Qed.
(* ********** *)
(* end of week-03_structural-induction-over-Peano-numbers.v *)