]> matita.cs.unibo.it Git - helm.git/commitdiff
some notation for map_arrows2
authorEnrico Tassi <enrico.tassi@inria.fr>
Tue, 6 Jul 2010 10:42:33 +0000 (10:42 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Tue, 6 Jul 2010 10:42:33 +0000 (10:42 +0000)
helm/software/matita/contribs/formal_topology/overlap/basic_pairs_to_o-basic_pairs.ma
helm/software/matita/contribs/formal_topology/overlap/basic_topologies.ma
helm/software/matita/contribs/formal_topology/overlap/categories.ma
helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma
helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs.ma
helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma

index 65709f4650c13c3941af2ebde7f700a24935f8d3..b72e4a5fae72d1df09708be619428ff342bd88f2 100644 (file)
@@ -19,9 +19,9 @@ include "relations_to_o-algebra.ma".
 definition o_basic_pair_of_basic_pair: basic_pair → Obasic_pair.
  intro b;
  constructor 1;
-  [ apply (map_objs2 ?? POW (concr b));
-  | apply (map_objs2 ?? POW (form b));
-  | apply (map_arrows2 ?? POW (concr b) (form b) (rel b)); ]
+  [ apply (POW (concr b));
+  | apply (POW (form b));
+  | apply (POW⎽⇒ ?); apply (rel b); ]
 qed.
 
 definition o_relation_pair_of_relation_pair:
@@ -29,7 +29,7 @@ definition o_relation_pair_of_relation_pair:
   Orelation_pair (o_basic_pair_of_basic_pair BP1) (o_basic_pair_of_basic_pair BP2).
  intros;
  constructor 1;
-  [ apply (map_arrows2 ?? POW (concr BP1) (concr BP2) (r \sub \c));
+  [ unfold o_basic_pair_of_basic_pair; simplify; apply (POW⎽⇒ ?); apply (r\sub \c); 
   | apply (map_arrows2 ?? POW (form BP1) (form BP2) (r \sub \f));
   | apply (.= (respects_comp2 ?? POW (concr BP1) (concr BP2) (form BP2)  r\sub\c (⊩\sub BP2) )^-1);
     cut ( ⊩ \sub BP2 ∘ r \sub \c =_12 r\sub\f ∘ ⊩ \sub BP1) as H;
@@ -115,7 +115,7 @@ qed.
 
 theorem BP_to_OBP_faithful:
  ∀S,T.∀f,g:arrows2 (category2_of_category1 BP) S T.
-  map_arrows2 ?? BP_to_OBP ?? f = map_arrows2 ?? BP_to_OBP ?? g → f=g.
+   BP_to_OBP⎽⇒ f = BP_to_OBP⎽⇒ g → f=g.
  intros; change with ( (⊩) ∘ f \sub \c = (⊩) ∘ g \sub \c);
  apply (POW_faithful);
  apply (.= respects_comp2 ?? POW (concr S) (concr T) (form T) f \sub \c (⊩ \sub T));
@@ -126,7 +126,7 @@ theorem BP_to_OBP_faithful:
 qed.
 
 theorem BP_to_OBP_full: 
-   ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? BP_to_OBP S T g = f).
+   ∀S,T.∀f. exT22 ? (λg:arrows2 ? S T. BP_to_OBP⎽⇒ g = f).
  intros; 
  cases (POW_full (concr S) (concr T) (Oconcr_rel ?? f)) (gc Hgc);
  cases (POW_full (form S) (form T) (Oform_rel ?? f)) (gf Hgf);
index b24a730cb05381873fc80103b8491ab1288e6f8e..5cb82832a167a8bafa07910d59b8c5e60e973a93 100644 (file)
@@ -145,9 +145,9 @@ definition BTop: category1.
             = minus_star_image o2 o3 b' (A o2 (minus_star_image o1 o2 a' (A o1 X))));
         [2: intro; apply sym1; apply (.= #‡(†((H' ?)\sup -1))); apply sym1; apply (K X);]
        clear K H' H1';
-       alias symbol "compose" (instance 2) = "category1 composition".
-cut (∀X:Ω \sup o1.
-              minus_star_image o1 o3 (b ∘ a) (A o1 X) = minus_star_image o1 o3 (b'∘a') (A o1 X));
+alias symbol "compose" (instance 1) = "category1 composition".
+cut (∀X:Ω^o1.
+              minus_star_image ?? (b ∘ a) (A o1 X) =_1 minus_star_image ?? (b'∘a') (A o1 X));
         [2: intro;
             apply (.= (minus_star_image_comp ??????));
             apply (.= #‡(saturated ?????));
index 72af642c34ae01fd03d058f197d3e0dad87d37d4..65320ae53efe492e1a29e9202076d1d1b6eb7c4e 100644 (file)
@@ -355,6 +355,10 @@ record functor2 (C1: category2) (C2: category2) : Type3 ≝
      ∀o1,o2,o3.∀f1:arrows2 ? o1 o2.∀f2:arrows2 ? o2 o3.
      map_arrows2 ?? (f2 ∘ f1) = map_arrows2 ?? f2 ∘ map_arrows2 ?? f1}.
 
+notation > "F⎽⇒ x" left associative with precedence 60 for @{'map_arrows2 $F $x}.
+notation "F\sub⇒ x" left associative with precedence 60 for @{'map_arrows2 $F $x}.
+interpretation "map_arrows2" 'map_arrows2 F x = (fun12 ?? (map_arrows2 ?? F ??) x).
+
 definition functor2_setoid: category2 → category2 → setoid3.
  intros (C1 C2);
  constructor 1;
index 915afc26d5229474d1a92f62bf7d2da959b41d6b..a86d286bc0b034d658e0b34521964795d29d540a 100644 (file)
@@ -324,6 +324,10 @@ coercion ORelation_setoid_of_arrows2_OA.
 
 prefer coercion Type_OF_objs2.
 
+notation > "B ⇒_\o2 C" right associative with precedence 72 for @{'arrows2_OA $B $C}.
+notation "B ⇒\sub (\o 2) C" right associative with precedence 72 for @{'arrows2_OA $B $C}.
+interpretation "'arrows2_OA" 'arrows2_OA A B = (arrows2 OA A B).
+
 (* qui la notazione non va *)
 lemma leq_to_eq_join: ∀S:OA.∀p,q:S. p ≤ q → q = (binary_join ? p q).
  intros;
index 6f16b0d3f05f9a0f0812f718e7a3c2066cbb6571..f0e0b712c80caefd132d18b220c6652ed706c975 100644 (file)
@@ -23,10 +23,6 @@ record Obasic_pair: Type2 ≝ {
 interpretation "o-basic pair relation indexed" 'Vdash2 x y c = (Orel c x y).
 interpretation "o-basic pair relation (non applied)" 'Vdash c = (Orel c).
 
-notation > "B ⇒_\o2 C" right associative with precedence 72 for @{'arrows2_OA $B $C}.
-notation "B ⇒\sub (\o 2) C" right associative with precedence 72 for @{'arrows2_OA $B $C}.
-interpretation "'arrows2_OA" 'arrows2_OA A B = (arrows2 OA A B).
-
 record Orelation_pair (BP1,BP2: Obasic_pair): Type2 ≝ { 
    Oconcr_rel: (Oconcr BP1) ⇒_\o2 (Oconcr BP2); Oform_rel: (Oform BP1) ⇒_\o2 (Oform BP2);
    Ocommute: ⊩ ∘ Oconcr_rel =_2 Oform_rel ∘ ⊩
@@ -248,4 +244,4 @@ interpretation "Universal pre-image ⊩*" 'rest x = (fun12 ? ? (or_f_star ? ?) (
 
 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 ? ?) (Orel x)).
\ No newline at end of file
+interpretation "Existential pre-image ⊩⎻" 'ext x = (fun12 ? ? (or_f_minus ? ?) (Orel x)).
index 9bcd51c7c0ece3c3acab8416844cf523e3f3df16..8bf57da98bfccfb58e524cd2e8219ae34d79965d 100644 (file)
@@ -51,7 +51,7 @@ qed.
 definition powerset_of_POW': ∀A.oa_P (POW' A) → Ω^A ≝ λA,x.x.
 coercion powerset_of_POW'.
 
-definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (POW' o1) (POW' o2).
+definition orelation_of_relation: ∀o1,o2:REL. o1 ⇒_\r1 o2 → (POW' o1) ⇒_\o2 (POW' o2).
  intros;
  constructor 1;
   [ constructor 1; 
@@ -90,7 +90,7 @@ definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (
 qed.
 
 lemma orelation_of_relation_preserves_equality:
- ∀o1,o2:REL.∀t,t': arrows1 ? o1 o2. 
+ ∀o1,o2:REL.∀t,t': o1 ⇒_\r1 o2. 
    t = t' → orelation_of_relation ?? t =_2 orelation_of_relation ?? t'.
  intros; split; unfold orelation_of_relation; simplify; intro; split; intro;
  simplify; whd in o1 o2;
@@ -140,7 +140,7 @@ qed.
 alias symbol "eq" = "setoid2 eq".
 alias symbol "compose" = "category1 composition".
 lemma orelation_of_relation_preserves_composition:
- ∀o1,o2,o3:REL.∀F: arrows1 ? o1 o2.∀G: arrows1 ? o2 o3.
+ ∀o1,o2,o3:REL.∀F: o1 ⇒_\r1 o2.∀G: o2 ⇒_\r1 o3.
   orelation_of_relation ?? (G ∘ F) = 
   comp2 OA ??? (orelation_of_relation ?? F) (orelation_of_relation ?? G).
  intros; split; intro; split; whd; intro; whd in ⊢ (% → %); intros;
@@ -172,7 +172,7 @@ qed.
 
 theorem POW_faithful:
  ∀S,T.∀f,g:arrows2 (category2_of_category1 REL) S T.
-  map_arrows2 ?? POW ?? f = map_arrows2 ?? POW ?? g → f = g.
+   POW⎽⇒ f =_2 POW⎽⇒ g → f =_2 g.
  intros; unfold POW in e; simplify in e; cases e;
  unfold orelation_of_relation in e3; simplify in e3; clear e e1 e2 e4;
  intros 2; cases (e3 {(x)}); 
@@ -186,13 +186,7 @@ lemma currify: ∀A,B,C. (A × B ⇒_1 C) → A → (B ⇒_1 C).
 intros; constructor 1; [ apply (b c); | intros; apply (#‡e);]
 qed.
 
-(*
-interpretation "lifting singl" 'singl x = 
- (fun11 ? (objs2 (POW ?))  (singleton ?) x).
-*)
-
-
-theorem POW_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? POW S T g = f).
+theorem POW_full: ∀S,T.∀f: (POW S) ⇒_\o2 (POW T) . exT22 ? (λg. POW⎽⇒ g = f).
  intros; exists;
   [ constructor 1; constructor 1;
      [ apply (λx:carr S.λy:carr T. y ∈ f {(x)});