]> matita.cs.unibo.it Git - helm.git/commitdiff
More properties of iff
authorAndrea Asperti <andrea.asperti@unibo.it>
Mon, 5 Dec 2011 08:56:30 +0000 (08:56 +0000)
committerAndrea Asperti <andrea.asperti@unibo.it>
Mon, 5 Dec 2011 08:56:30 +0000 (08:56 +0000)
matita/matita/lib/basics/logic.ma

index c80fa31f55c5469a97212e3c49431360e3e805fa..8f981d46924556ee885407a1320d147ef07aaad0 100644 (file)
@@ -153,7 +153,28 @@ inductive ex2 (A:Type[0]) (P,Q:A →Prop) : Prop ≝
 definition iff :=
  λ A,B. (A → B) ∧ (B → A).
 
-interpretation "iff" 'iff a b = (iff a b).  
+interpretation "iff" 'iff a b = (iff a b).
+
+lemma iff_sym: ∀A,B. A ↔ B → B ↔ A.
+#A #B * /3/ qed.
+
+lemma iff_trans:∀A,B,C. A ↔ B → B ↔ C → A ↔ C.
+#A #B #C * #H1 #H2 * #H3 #H4 % /3/ qed.
+
+lemma iff_not: ∀A,B. A ↔ B → ¬A ↔ ¬B.
+#A #B * #H1 #H2 % /3/ qed.
+
+lemma iff_and_l: ∀A,B,C. A ↔ B → C ∧ A ↔ C ∧ B.
+#A #B #C * #H1 #H2 % * /3/ qed.  
+
+lemma iff_and_r: ∀A,B,C. A ↔ B → A ∧ C ↔ B ∧ C.
+#A #B #C * #H1 #H2 % * /3/ qed.  
+
+lemma iff_or_l: ∀A,B,C. A ↔ B → C ∨ A ↔ C ∨ B.
+#A #B #C * #H1 #H2 % * /3/ qed.  
+
+lemma iff_or_r: ∀A,B,C. A ↔ B → A ∨ C ↔ B ∨ C.
+#A #B #C * #H1 #H2 % * /3/ qed.  
 
 (* cose per destruct: da rivedere *)