aboutsummaryrefslogtreecommitdiffhomepage
path: root/theories/Logic/Decidable.v
diff options
context:
space:
mode:
authorGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2008-02-02 15:51:00 +0000
committerGravatar letouzey <letouzey@85f007b7-540e-0410-9357-904b9bb8a0f7>2008-02-02 15:51:00 +0000
commit456e7bb78ab46ccd5ea0c34726356c7c014308d8 (patch)
tree59164f7e906f24e85c4a3d345cd94d4f5ffddb80 /theories/Logic/Decidable.v
parentc61b48f8b123e572b33c6d080a2b09aa5ecde979 (diff)
factorization part II (Properties + EqProperties), inclusion of FSetDecide (from A. Bohannon)
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@10500 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'theories/Logic/Decidable.v')
-rw-r--r--theories/Logic/Decidable.v159
1 files changed, 147 insertions, 12 deletions
diff --git a/theories/Logic/Decidable.v b/theories/Logic/Decidable.v
index 1995ca77a..6f793ce2f 100644
--- a/theories/Logic/Decidable.v
+++ b/theories/Logic/Decidable.v
@@ -12,49 +12,184 @@
Definition decidable (P:Prop) := P \/ ~ P.
Theorem dec_not_not : forall P:Prop, decidable P -> (~ P -> False) -> P.
-unfold decidable in |- *; tauto.
+Proof.
+unfold decidable; tauto.
Qed.
Theorem dec_True : decidable True.
-unfold decidable in |- *; auto.
+Proof.
+unfold decidable; auto.
Qed.
Theorem dec_False : decidable False.
-unfold decidable, not in |- *; auto.
+Proof.
+unfold decidable, not; auto.
Qed.
Theorem dec_or :
forall A B:Prop, decidable A -> decidable B -> decidable (A \/ B).
-unfold decidable in |- *; tauto.
+Proof.
+unfold decidable; tauto.
Qed.
Theorem dec_and :
forall A B:Prop, decidable A -> decidable B -> decidable (A /\ B).
-unfold decidable in |- *; tauto.
+Proof.
+unfold decidable; tauto.
Qed.
Theorem dec_not : forall A:Prop, decidable A -> decidable (~ A).
-unfold decidable in |- *; tauto.
+Proof.
+unfold decidable; tauto.
Qed.
Theorem dec_imp :
forall A B:Prop, decidable A -> decidable B -> decidable (A -> B).
-unfold decidable in |- *; tauto.
+Proof.
+unfold decidable; tauto.
+Qed.
+
+Theorem dec_iff :
+ forall A B:Prop, decidable A -> decidable B -> decidable (A<->B).
+Proof.
+unfold decidable; tauto.
Qed.
Theorem not_not : forall P:Prop, decidable P -> ~ ~ P -> P.
-unfold decidable in |- *; tauto. Qed.
+Proof.
+unfold decidable; tauto.
+Qed.
Theorem not_or : forall A B:Prop, ~ (A \/ B) -> ~ A /\ ~ B.
-tauto. Qed.
+Proof.
+tauto.
+Qed.
Theorem not_and : forall A B:Prop, decidable A -> ~ (A /\ B) -> ~ A \/ ~ B.
-unfold decidable in |- *; tauto. Qed.
+Proof.
+unfold decidable; tauto.
+Qed.
Theorem not_imp : forall A B:Prop, decidable A -> ~ (A -> B) -> A /\ ~ B.
-unfold decidable in |- *; tauto.
+Proof.
+unfold decidable; tauto.
Qed.
Theorem imp_simp : forall A B:Prop, decidable A -> (A -> B) -> ~ A \/ B.
-unfold decidable in |- *; tauto.
+Proof.
+unfold decidable; tauto.
+Qed.
+
+(** Results formulated with iff, used in FSetDecide.
+ Negation are expanded since it is unclear whether setoid rewrite
+ will always perform conversion. *)
+
+(** We begin with lemmas that, when read from left to right,
+ can be understood as ways to eliminate uses of [not]. *)
+
+Theorem not_true_iff : (True -> False) <-> False.
+Proof.
+tauto.
+Qed.
+
+Theorem not_false_iff : (False -> False) <-> True.
+Proof.
+tauto.
+Qed.
+
+Theorem not_not_iff : forall A:Prop, decidable A ->
+ (((A -> False) -> False) <-> A).
+Proof.
+unfold decidable; tauto.
+Qed.
+
+Theorem contrapositive : forall A B:Prop, decidable A ->
+ (((A -> False) -> (B -> False)) <-> (B -> A)).
+Proof.
+unfold decidable; tauto.
+Qed.
+
+Lemma or_not_l_iff_1 : forall A B: Prop, decidable A ->
+ ((A -> False) \/ B <-> (A -> B)).
+Proof.
+unfold decidable. tauto.
+Qed.
+
+Lemma or_not_l_iff_2 : forall A B: Prop, decidable B ->
+ ((A -> False) \/ B <-> (A -> B)).
+Proof.
+unfold decidable. tauto.
+Qed.
+
+Lemma or_not_r_iff_1 : forall A B: Prop, decidable A ->
+ (A \/ (B -> False) <-> (B -> A)).
+Proof.
+unfold decidable. tauto.
Qed.
+
+Lemma or_not_r_iff_2 : forall A B: Prop, decidable B ->
+ (A \/ (B -> False) <-> (B -> A)).
+Proof.
+unfold decidable. tauto.
+Qed.
+
+Lemma imp_not_l : forall A B: Prop, decidable A ->
+ (((A -> False) -> B) <-> (A \/ B)).
+Proof.
+unfold decidable. tauto.
+Qed.
+
+
+(** Moving Negations Around:
+ We have four lemmas that, when read from left to right,
+ describe how to push negations toward the leaves of a
+ proposition and, when read from right to left, describe
+ how to pull negations toward the top of a proposition. *)
+
+Theorem not_or_iff : forall A B:Prop,
+ (A \/ B -> False) <-> (A -> False) /\ (B -> False).
+Proof.
+tauto.
+Qed.
+
+Lemma not_and_iff : forall A B:Prop,
+ (A /\ B -> False) <-> (A -> B -> False).
+Proof.
+tauto.
+Qed.
+
+Lemma not_imp_iff : forall A B:Prop, decidable A ->
+ (((A -> B) -> False) <-> A /\ (B -> False)).
+Proof.
+unfold decidable. tauto.
+Qed.
+
+Lemma not_imp_rev_iff : forall A B : Prop, decidable A ->
+ (((A -> B) -> False) <-> (B -> False) /\ A).
+Proof.
+unfold decidable. tauto.
+Qed.
+
+
+
+(** With the following hint database, we can leverage [auto] to check
+ decidability of propositions. *)
+
+Hint Resolve dec_True dec_False dec_or dec_and dec_imp dec_not dec_iff
+ : decidable_prop.
+
+(** [solve_decidable using lib] will solve goals about the
+ decidability of a proposition, assisted by an auxiliary
+ database of lemmas. The database is intended to contain
+ lemmas stating the decidability of base propositions,
+ (e.g., the decidability of equality on a particular
+ inductive type). *)
+
+Tactic Notation "solve_decidable" "using" ident(db) :=
+ match goal with
+ | |- decidable _ =>
+ solve [ auto 100 with decidable_prop db ]
+ end.
+
+Tactic Notation "solve_decidable" :=
+ solve_decidable using core.