]> matita.cs.unibo.it Git - helm.git/commitdiff
Added injective compose
authorWilmer Ricciotti <ricciott@cs.unibo.it>
Sun, 21 Jun 2009 18:32:22 +0000 (18:32 +0000)
committerWilmer Ricciotti <ricciott@cs.unibo.it>
Sun, 21 Jun 2009 18:32:22 +0000 (18:32 +0000)
helm/software/matita/library/higher_order_defs/functions.ma

index e14b38db6b8c6ff15d8cb15c548f33e6b42b6c36..90f550618d9e39e047c0099726ed5b6613c89964 100644 (file)
@@ -58,3 +58,7 @@ definition distributive2: \forall A,B:Type.\forall f:A \to B \to B.
 \forall g: B\to B\to B. Prop
 \def \lambda A,B.\lambda f,g.\forall x:A.\forall y,z:B. f x (g y z) = g (f x y) (f x z).
 
+lemma injective_compose : ∀A,B,C,f,g.injective A B f → injective B C g → 
+  injective A C (λx.g (f x)).
+intros;unfold;intros;simplify in H2;apply H;apply H1;assumption;
+qed.
\ No newline at end of file