X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Fbasic_pairs.ma;h=0734411f8d07192881bcb29516b089959d80f2d8;hb=11a22c74b3b2307eedf89c0439ba02d199dcdc9e;hp=c5546477b938831b5cd22363de05a5997b10686e;hpb=23043db144b24b8cd2072800b61137bb396f891e;p=helm.git diff --git a/helm/software/matita/contribs/formal_topology/overlap/basic_pairs.ma b/helm/software/matita/contribs/formal_topology/overlap/basic_pairs.ma index c5546477b..0734411f8 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/basic_pairs.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/basic_pairs.ma @@ -15,42 +15,26 @@ include "relations.ma". include "notation.ma". -record basic_pair: Type1 ≝ - { concr: REL; - form: REL; - rel: arrows1 ? concr form - }. +record basic_pair: Type1 ≝ { + concr: REL; form: REL; rel: concr ⇒_\r1 form +}. -interpretation "basic pair relation" 'Vdash2 x y c = (fun21 ___ (rel c) x y). +interpretation "basic pair relation" 'Vdash2 x y c = (fun21 ??? (rel c) x y). interpretation "basic pair relation (non applied)" 'Vdash c = (rel c). -alias symbol "eq" = "setoid1 eq". -alias symbol "compose" = "category1 composition". -record relation_pair (BP1,BP2: basic_pair): Type1 ≝ - { concr_rel: arrows1 ? (concr BP1) (concr BP2); - form_rel: arrows1 ? (form BP1) (form BP2); - commute: ⊩ ∘ concr_rel = form_rel ∘ ⊩ +record relation_pair (BP1,BP2: basic_pair): Type1 ≝ { + concr_rel: (concr BP1) ⇒_\r1 (concr BP2); form_rel: (form BP1) ⇒_\r1 (form BP2); + commute: ⊩ ∘ concr_rel =_1 form_rel ∘ ⊩ }. +interpretation "concrete relation" 'concr_rel r = (concr_rel ?? r). +interpretation "formal relation" 'form_rel r = (form_rel ?? r). -interpretation "concrete relation" 'concr_rel r = (concr_rel __ r). -interpretation "formal relation" 'form_rel r = (form_rel __ r). - -definition relation_pair_equality: - ∀o1,o2. equivalence_relation1 (relation_pair o1 o2). - intros; - constructor 1; - [ apply (λr,r'. ⊩ ∘ r \sub\c = ⊩ ∘ r' \sub\c); - | simplify; - intros; - apply refl1; - | simplify; - intros 2; - apply sym1; - | simplify; - intros 3; - apply trans1; - ] +definition relation_pair_equality: ∀o1,o2. equivalence_relation1 (relation_pair o1 o2). + intros; constructor 1; [ apply (λr,r'. ⊩ ∘ r \sub\c = ⊩ ∘ r' \sub\c); + | simplify; intros; apply refl1; + | simplify; intros 2; apply sym1; + | simplify; intros 3; apply trans1; ] qed. definition relation_pair_setoid: basic_pair → basic_pair → setoid1. @@ -66,7 +50,7 @@ definition relation_pair_of_relation_pair_setoid : coercion relation_pair_of_relation_pair_setoid. lemma eq_to_eq': - ∀o1,o2.∀r,r':relation_pair_setoid o1 o2. r=r' → r \sub\f ∘ ⊩ = r'\sub\f ∘ ⊩. + ∀o1,o2.∀r,r':relation_pair_setoid o1 o2. r =_1 r' → r \sub\f ∘ ⊩ = r'\sub\f ∘ ⊩. intros 7 (o1 o2 r r' H c1 f2); split; intro H1; [ lapply (fi ?? (commute ?? r c1 f2) H1) as H2; @@ -192,7 +176,7 @@ definition relation_pair_setoid_of_arrows1_BP : ∀P,Q. arrows1 BP P Q → relation_pair_setoid P Q ≝ λP,Q,x.x. coercion relation_pair_setoid_of_arrows1_BP. -definition BPext: ∀o: BP. form o ⇒ Ω \sup (concr o). +definition BPext: ∀o: BP. (form o) ⇒_1 Ω^(concr o). intros; constructor 1; [ apply (ext ? ? (rel o)); | intros; @@ -200,13 +184,13 @@ definition BPext: ∀o: BP. form o ⇒ Ω \sup (concr o). apply refl1] qed. -definition BPextS: ∀o: BP. Ω \sup (form o) ⇒ Ω \sup (concr o). +definition BPextS: ∀o: BP. Ω^(form o) ⇒_1 Ω^(concr o). intros; constructor 1; [ apply (minus_image ?? (rel o)); | intros; apply (#‡e); ] qed. -definition fintersects: ∀o: BP. binary_morphism1 (form o) (form o) (Ω \sup (form o)). +definition fintersects: ∀o: BP. (form o) × (form o) ⇒_1 Ω^(form o). intros (o); constructor 1; [ apply (λa,b: form o.{c | BPext o c ⊆ BPext o a ∩ BPext o b }); intros; simplify; apply (.= (†e)‡#); apply refl1 @@ -215,27 +199,27 @@ definition fintersects: ∀o: BP. binary_morphism1 (form o) (form o) (Ω \sup (f | apply (. #‡((†e)‡(†e1))); assumption]] qed. -interpretation "fintersects" 'fintersects U V = (fun21 ___ (fintersects _) U V). +interpretation "fintersects" 'fintersects U V = (fun21 ??? (fintersects ?) U V). definition fintersectsS: - ∀o:BP. binary_morphism1 (Ω \sup (form o)) (Ω \sup (form o)) (Ω \sup (form o)). + ∀o:BP. Ω^(form o) × Ω^(form o) ⇒_1 Ω^(form o). intros (o); constructor 1; - [ apply (λo: basic_pair.λa,b: Ω \sup (form o).{c | BPext o c ⊆ BPextS o a ∩ BPextS o b }); + [ apply (λo: basic_pair.λa,b: Ω^(form o).{c | BPext o c ⊆ BPextS o a ∩ BPextS o b }); intros; simplify; apply (.= (†e)‡#); apply refl1 | intros; split; simplify; intros; [ apply (. #‡((†e^-1)‡(†e1^-1))); assumption | apply (. #‡((†e)‡(†e1))); assumption]] qed. -interpretation "fintersectsS" 'fintersects U V = (fun21 ___ (fintersectsS _) U V). +interpretation "fintersectsS" 'fintersects U V = (fun21 ??? (fintersectsS ?) U V). -definition relS: ∀o: BP. binary_morphism1 (concr o) (Ω \sup (form o)) CPROP. +definition relS: ∀o: BP. (concr o) × Ω^(form o) ⇒_1 CPROP. intros (o); constructor 1; - [ apply (λx:concr o.λS: Ω \sup (form o).∃y:form o.y ∈ S ∧ x ⊩_o y); + [ apply (λx:concr o.λS: Ω^(form o).∃y:form o.y ∈ S ∧ x ⊩⎽o y); | intros; split; intros; cases e2; exists [1,3: apply w] [ apply (. (#‡e1^-1)‡(e^-1‡#)); assumption | apply (. (#‡e1)‡(e‡#)); assumption]] qed. -interpretation "basic pair relation for subsets" 'Vdash2 x y c = (fun21 (concr _) __ (relS c) x y). -interpretation "basic pair relation for subsets (non applied)" 'Vdash c = (fun21 ___ (relS c)). +interpretation "basic pair relation for subsets" 'Vdash2 x y c = (fun21 (concr ?) ?? (relS c) x y). +interpretation "basic pair relation for subsets (non applied)" 'Vdash c = (fun21 ??? (relS c)).