X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Frelations_to_o-algebra.ma;fp=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Frelations_to_o-algebra.ma;h=7f327c4972130fcb143b5aa1eb45f2eafb6bbc60;hb=d87c05ef2d706cc9b69454c191568028134138ea;hp=a7041357af3a0cffd16f257a87008e4be21aac7f;hpb=4bdb34a1cce33b4387b04cc37bf229e08f5bbafb;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 a7041357a..7f327c497 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 @@ -169,5 +169,14 @@ definition SUBSETS': carr3 (arrows3 CAT2 (category2_of_category1 REL) OA). apply (.= (orelation_of_relation_preserves_composition o1 o2 o4 f1 (f3∘f2))); apply (#‡(orelation_of_relation_preserves_composition o2 o3 o4 f2 f3)); ] qed. - - \ No newline at end of file + +theorem SUBSETS_faithful: + ∀S,T.∀f,g:arrows2 (category2_of_category1 REL) S T. + map_arrows2 ?? SUBSETS' ?? f = map_arrows2 ?? SUBSETS' ?? g → f=g. + intros; unfold SUBSETS' in e; simplify in e; cases e; + unfold orelation_of_relation in e3; simplify in e3; clear e e1 e2 e4; + intros 2; lapply (e3 (singleton ? x)); cases Hletin; + split; intro; [ lapply (s y); | lapply (s1 y); ] + [2,4: exists; [1,3:apply x] split; [1,3: assumption |*: change with (x=x); apply rule #] + |*: cases Hletin1; cases x1; change in f3 with (eq1 ? x w); apply (. f3‡#); assumption;] +qed. \ No newline at end of file