X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Frelations_to_o-algebra.ma;h=bc13ecfd7ee5ee6e259630e5a1a18a8c0de76a26;hb=ddc0a7b3f0acd57f879e540e696f69ca0c20bbf5;hp=af3015eb368737d5f62f3ce5c8689399edfc9e07;hpb=4735518bc617c052344feefa75690b613041a025;p=helm.git diff --git a/helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma b/helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma index af3015eb3..bc13ecfd7 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma @@ -182,6 +182,23 @@ theorem POW_faithful: |*: cases Hletin; cases x1; change in f3 with (x =_1 w); apply (. f3‡#); assumption;] qed. + +lemma currify: ∀A,B,C. binary_morphism1 A B C → A → unary_morphism1 B C. +intros; constructor 1; [ apply (b c); | intros; apply (#‡e);] +qed. + +(* +alias symbol "singl" = "singleton". +alias symbol "eq" = "setoid eq". +lemma in_singleton_to_eq : ∀A:setoid.∀y,x:A.y ∈ {(x)} → (eq1 A) y x. +intros; apply sym1; apply f; +qed. + +lemma eq_to_in_singleton : ∀A:setoid.∀y,x:A.eq1 A y x → y ∈ {(x)}. +intros; apply (e^-1); +qed. +*) + interpretation "lifting singl" 'singl x = (fun11 _ (objs2 (POW _)) (singleton _) x).