]> matita.cs.unibo.it Git - helm.git/commitdiff
a bit of work done while travelling to padova
authorEnrico Tassi <enrico.tassi@inria.fr>
Wed, 21 Jan 2009 13:57:40 +0000 (13:57 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Wed, 21 Jan 2009 13:57:40 +0000 (13:57 +0000)
helm/software/matita/contribs/formal_topology/overlap/basic_pairs_to_o-basic_pairs.ma
helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs.ma

index 7dbd3a80ae6bd0c9bbe9b6c687e9404dda6420d8..2aec4f7e8e897300c28e6c2c142e521a91179c6c 100644 (file)
@@ -79,23 +79,34 @@ definition BP_to_OBP: carr3 (arrows3 CAT2 (category2_of_category1 BP) OBP).
     apply (respects_comp2 ?? SUBSETS' (concr o1) (concr o2) (concr o3) f1\sub\c f2\sub\c);]     
 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.
- intros; unfold BP_to_OBP in e; simplify in e; cases e;
unfold orelation_of_relation in e3; simplify in e3; clear e e1 e2 e4;
intros 2; change in match or_f_ in e3 with (λq,w,x.fun12 ?? (or_f q w) x);
simplify in e3; STOP lapply (e3 (singleton ? x)); cases Hletin;
- split; intro; [ lapply (s y); | lapply (s1 y); ]
-  [2,4: exists; [1,3:apply x] split; [1,3: assumption |*: change with (x=x); apply rule #]
-  |*: cases Hletin1; cases x1; change in f3 with (eq1 ? x w); apply (. f3‡#); assumption;]
+ intros; change with ( (⊩) ∘ f \sub \c = (⊩) ∘ g \sub \c);
apply (SUBSETS_faithful);
apply (.= respects_comp2 ?? SUBSETS' (concr S) (concr T) (form T) f \sub \c (⊩ \sub T));
apply sym2;
+ apply (.= respects_comp2 ?? SUBSETS' (concr S) (concr T) (form T) g \sub \c (⊩ \sub T));
+ apply sym2;
+ apply e;
 qed.
-*)
 
-(*
-theorem SUBSETS_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? BP_to_OBP S T g = f).
- intros; exists;
-*)
\ No newline at end of file
+theorem BP_to_OBP_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? BP_to_OBP S T g = f).
+ intros; 
+ cases (SUBSETS_full (concr S) (concr T) (Oconcr_rel ?? f)) (gc Hgc);
+ cases (SUBSETS_full (form S) (form T) (Oform_rel ?? f)) (gf Hgf);
+ exists[
+   constructor 1; [apply gc|apply gf]
+   apply (SUBSETS_faithful);
+   apply (let xxxx ≝SUBSETS' in .= respects_comp2 ?? SUBSETS' (concr S) (concr T) (form T) gc (rel T));
+   apply rule (.= Hgc‡#);
+   apply (.= Ocommute ?? f);
+   apply (.= #‡Hgf^-1);
+   apply (let xxxx ≝SUBSETS' in (respects_comp2 ?? SUBSETS' (concr S) (form S) (form T) (rel S) gf)^-1)]
+ split;
+  [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x); 
+  | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
+  | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
+  | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
+ simplify; apply (†(Hgc‡#));
+qed.   
index 6476c4aebf727ac77847eac2ea097fa045484c25..b6d6e1b3dd23cb088149fb1c509e25a0fd050fc8 100644 (file)
@@ -21,6 +21,7 @@ record Obasic_pair: Type2 ≝
    Orel: arrows2 ? Oconcr Oform
  }.
 
+(* FIX *)
 interpretation "basic pair relation indexed" 'Vdash2 x y c = (Orel c x y).
 interpretation "basic pair relation (non applied)" 'Vdash c = (Orel c).
 
@@ -35,7 +36,8 @@ record Orelation_pair (BP1,BP2: Obasic_pair): Type2 ≝
    Oform_rel: arrows2 ? (Oform BP1) (Oform BP2);
    Ocommute: ⊩ ∘ Oconcr_rel = Oform_rel ∘ ⊩
  }.
-
+(* FIX *)
 interpretation "concrete relation" 'concr_rel r = (Oconcr_rel __ r). 
 interpretation "formal relation" 'form_rel r = (Oform_rel __ r).