nus/cs3234/labs/week-07_soundness-and-completeness-of-equality-predicates-revisited.v
2024-05-08 16:20:16 +08:00

282 lines
5.6 KiB
Coq

(* week-07_soundness-and-completeness-of-equality-predicates-revisited.v *)
(* LPP 2024 - S3C234 2023-2024, Sem2 *)
(* Olivier Danvy <danvy@yale-nus.edu.sg> *)
(* 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 *)