]> matita.cs.unibo.it Git - helm.git/commitdiff
not_to_not
authorAndrea Asperti <andrea.asperti@unibo.it>
Wed, 3 Jun 2009 14:42:12 +0000 (14:42 +0000)
committerAndrea Asperti <andrea.asperti@unibo.it>
Wed, 3 Jun 2009 14:42:12 +0000 (14:42 +0000)
helm/software/matita/library/logic/connectives.ma

index fcd046277fda5e50dfdd8bce702d065b14b789a1..76d0323166ad947578481e7c06bb4d8d8b3cc41a 100644 (file)
@@ -32,6 +32,12 @@ qed.
 
 default "absurd" cic:/matita/logic/connectives/absurd.con.
 
+theorem not_to_not : \forall A,B:Prop. (A → B) \to ¬B →¬A.
+intros.unfold.intro.apply H1.apply (H H2).
+qed.
+
+default "absurd" cic:/matita/logic/connectives/absurd.con.
+
 inductive And (A,B:Prop) : Prop \def
     conj : A \to B \to (And A B).