X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Fo-basic_pairs.ma;h=6517689a15696f000993fded2ecf2076aff73dde;hb=fc577dad1529b2d90c40dad8e6e3429281107c99;hp=c5f125ac679dc3d66c704874d27736a87fa55a67;hpb=bb0fff7ebc68535a75e260082b7db26c1d99f643;p=helm.git diff --git a/helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs.ma b/helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs.ma index c5f125ac6..6517689a1 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs.ma @@ -20,16 +20,15 @@ record basic_pair: Type2 ≝ rel: arrows2 ? concr form }. -notation > "x ⊩ y" with precedence 45 for @{'Vdash2 $x $y ?}. -notation < "x (⊩ \below c) y" with precedence 45 for @{'Vdash2 $x $y $c}. -notation < "⊩ \sub c" with precedence 60 for @{'Vdash $c}. -notation > "⊩ " with precedence 60 for @{'Vdash ?}. - interpretation "basic pair relation indexed" 'Vdash2 x y c = (rel c x y). interpretation "basic pair relation (non applied)" 'Vdash c = (rel c). alias symbol "eq" = "setoid1 eq". alias symbol "compose" = "category1 composition". +(*DIFFER*) + +alias symbol "eq" = "setoid2 eq". +alias symbol "compose" = "category2 composition". record relation_pair (BP1,BP2: basic_pair): Type2 ≝ { concr_rel: arrows2 ? (concr BP1) (concr BP2); form_rel: arrows2 ? (form BP1) (form BP2); @@ -68,6 +67,10 @@ definition relation_pair_setoid: basic_pair → basic_pair → setoid2. ] qed. +definition relation_pair_of_relation_pair_setoid: + ∀P,Q. relation_pair_setoid P Q → relation_pair P Q ≝ λP,Q,x.x. +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 ∘ ⊩. intros 5 (o1 o2 r r' H); change in H with (⊩ ∘ r\sub\c = ⊩ ∘ r'\sub\c); apply (.= ((commute ?? r) \sup -1)); @@ -97,23 +100,23 @@ definition relation_pair_composition: | apply (r1 \sub\f ∘ r \sub\f) | lapply (commute ?? r) as H; lapply (commute ?? r1) as H1; - apply rule (.= ASSOC1); + apply rule (.= ASSOC); apply (.= #‡H1); - apply rule (.= ASSOC1\sup -1); + apply rule (.= ASSOC ^ -1); apply (.= H‡#); - apply rule ASSOC1] + apply rule ASSOC] | intros; change with (⊩ ∘ (b\sub\c ∘ a\sub\c) = ⊩ ∘ (b'\sub\c ∘ a'\sub\c)); change in e with (⊩ ∘ a \sub\c = ⊩ ∘ a' \sub\c); change in e1 with (⊩ ∘ b \sub\c = ⊩ ∘ b' \sub\c); - apply rule (.= ASSOC1); + apply rule (.= ASSOC); apply (.= #‡e1); apply (.= #‡(commute ?? b')); - apply rule (.= ASSOC1 \sup -1); + apply rule (.= ASSOC \sup -1); apply (.= e‡#); - apply rule (.= ASSOC1); + apply rule (.= ASSOC); apply (.= #‡(commute ?? b')\sup -1); - apply rule (ASSOC1 \sup -1)] + apply rule (ASSOC \sup -1)] qed. definition BP: category2. @@ -125,7 +128,7 @@ definition BP: category2. | intros; change with (⊩ ∘ (a34\sub\c ∘ (a23\sub\c ∘ a12\sub\c)) = ⊩ ∘ ((a34\sub\c ∘ a23\sub\c) ∘ a12\sub\c)); - apply rule (ASSOC1‡#); + apply rule (ASSOC‡#); | intros; change with (⊩ ∘ (a\sub\c ∘ (id_relation_pair o1)\sub\c) = ⊩ ∘ a\sub\c); apply ((id_neutral_right2 ????)‡#); @@ -134,6 +137,12 @@ definition BP: category2. apply ((id_neutral_left2 ????)‡#);] qed. +definition basic_pair_of_objs2_BP: objs2 BP → basic_pair ≝ λx.x. +coercion basic_pair_of_objs2_BP. + +definition relation_pair_setoid_of_arrows2_BP: + ∀P,Q.arrows2 BP P Q → relation_pair_setoid P Q ≝ λP,Q,c.c. +coercion relation_pair_setoid_of_arrows2_BP. (* definition BPext: ∀o: BP. form o ⇒ Ω \sup (concr o). @@ -184,4 +193,20 @@ qed. interpretation "basic pair relation for subsets" 'Vdash2 x y = (fun1 (concr _) __ (relS _) x y). interpretation "basic pair relation for subsets (non applied)" 'Vdash = (fun1 ___ (relS _)). -*) \ No newline at end of file +*) + +notation "□ \sub b" non associative with precedence 90 for @{'box $b}. +notation > "□_term 90 b" non associative with precedence 90 for @{'box $b}. +interpretation "Universal image ⊩⎻*" 'box x = (fun12 _ _ (or_f_minus_star _ _) (rel x)). + +notation "◊ \sub b" non associative with precedence 90 for @{'diamond $b}. +notation > "◊_term 90 b" non associative with precedence 90 for @{'diamond $b}. +interpretation "Existential image ⊩" 'diamond x = (fun12 _ _ (or_f _ _) (rel x)). + +notation "'Rest' \sub b" non associative with precedence 90 for @{'rest $b}. +notation > "'Rest'⎽term 90 b" non associative with precedence 90 for @{'rest $b}. +interpretation "Universal pre-image ⊩*" 'rest x = (fun12 _ _ (or_f_star _ _) (rel x)). + +notation "'Ext' \sub b" non associative with precedence 90 for @{'ext $b}. +notation > "'Ext'⎽term 90 b" non associative with precedence 90 for @{'ext $b}. +interpretation "Existential pre-image ⊩⎻" 'ext x = (fun12 _ _ (or_f_minus _ _) (rel x)).