From d87c05ef2d706cc9b69454c191568028134138ea Mon Sep 17 00:00:00 2001 From: Claudio Sacerdoti Coen Date: Sat, 17 Jan 2009 20:10:36 +0000 Subject: [PATCH] faithful --- .../overlap/relations_to_o-algebra.ma | 13 +++++++++++-- 1 file changed, 11 insertions(+), 2 deletions(-) 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 -- 2.39.2