]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/higher_order_defs/functions.ma
Added injective compose
[helm.git] / helm / software / matita / library / higher_order_defs / functions.ma
index a6174f48fa1b41dae53b4de72dc318f5b62fa7a1..90f550618d9e39e047c0099726ed5b6613c89964 100644 (file)
@@ -18,12 +18,7 @@ definition compose \def
   \lambda A,B,C:Type.\lambda f:(B\to C).\lambda g:(A\to B).\lambda x:A.
   f (g x).
 
-notation "hvbox(a break \circ b)" 
-  left associative with precedence 70
-for @{ 'compose $a $b }.
-
-interpretation "function composition" 'compose f g =
-  (cic:/matita/higher_order_defs/functions/compose.con _ _ _ f g).
+interpretation "function composition" 'compose f g = (compose ? ? ? f g).
 
 definition injective: \forall A,B:Type.\forall f:A \to B.Prop
 \def \lambda A,B.\lambda f.
@@ -63,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