]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs.ma
the new coercion behaviour (variants + composition with ID) and the new
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_pairs.ma
index 956a26af952c6f317741f4757f59609b6570dacb..6517689a15696f000993fded2ecf2076aff73dde 100644 (file)
@@ -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));
@@ -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).