X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fhigher_order_defs%2Ffunctions.ma;h=90f550618d9e39e047c0099726ed5b6613c89964;hb=7002fb8d9d0102e9baa410935fdabc9be0f8690d;hp=195cbfc171be63395d565d6f6adde9566fd25548;hpb=80ea6f314e89d9d280338c41860cb04949319629;p=helm.git diff --git a/helm/software/matita/library/higher_order_defs/functions.ma b/helm/software/matita/library/higher_order_defs/functions.ma index 195cbfc17..90f550618 100644 --- a/helm/software/matita/library/higher_order_defs/functions.ma +++ b/helm/software/matita/library/higher_order_defs/functions.ma @@ -18,8 +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). -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. @@ -59,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