X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fformal_topology%2Frelations.ma;h=3516999008aa714be19b99bf19b068dc585881dc;hb=cfccf434a57e10848d74d06674af4ec9cef0f0ca;hp=89273de051a3982f3825f27355f5d7ac0b52bc01;hpb=bc477154d15f6979c948f9af602199af547d193e;p=helm.git diff --git a/matita/matita/lib/formal_topology/relations.ma b/matita/matita/lib/formal_topology/relations.ma index 89273de05..351699900 100644 --- a/matita/matita/lib/formal_topology/relations.ma +++ b/matita/matita/lib/formal_topology/relations.ma @@ -12,8 +12,9 @@ (* *) (**************************************************************************) +include "basics/core_notation/comprehension_2.ma". include "formal_topology/subsets.ma". - +(* record binary_relation (A,B: SET) : Type[1] ≝ { satisfy:> binary_morphism1 A B CPROP }. @@ -115,7 +116,7 @@ interpretation "'arrows2_REL" 'arrows2_REL A B = (arrows2 (category2_of_category definition full_subset: ∀s:REL. Ω^s. - apply (λs.{x | ⊤}); + apply (λs.{x | True}); intros; simplify; split; intro; assumption. qed. @@ -316,9 +317,12 @@ theorem ext_comp: qed. *) +axiom daemon : False. + theorem extS_singleton: ∀o1,o2.∀a.∀x.extS o1 o2 a {(x)} = ext o1 o2 a x. intros; unfold extS; unfold ext; unfold singleton; simplify; split; intros 2; simplify; simplify in f; [ cases f; cases e; cases x1; change in f2 with (x =_1 w); apply (. #‡f2); assumption; | split; try apply I; exists [apply x] split; try assumption; change with (x = x); apply rule #;] qed. +*)