]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma
More re-organization.
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-concrete_spaces.ma
index 0bd8715a67550692dbc7fbdc702e8ecbdff15f86..78ea321a59857682393a42ed2baca1a4aa6fe0ce 100644 (file)
@@ -17,51 +17,44 @@ include "o-saturations.ma".
 
 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 = (fun_1 _ _ (or_f_minus_star _ _) (rel x)).
+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 = (fun_1 _ _ (or_f _ _) (rel x)).
+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 = (fun_1 _ _ (or_f_star _ _) (rel x)).
+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 = (fun_1 _ _ (or_f_minus _ _) (rel x)).
+interpretation "Existential pre-image ⊩⎻" 'ext x = (fun12 _ _ (or_f_minus _ _) (rel x)).
 
-lemma hint : ∀p,q.arrows1 OA p q → ORelation_setoid p q.
-intros; assumption;
-qed.
-
-coercion hint nocomposites.
-
-definition A : ∀b:BP. unary_morphism (oa_P (form b)) (oa_P (form b)).
+definition A : ∀b:BP. unary_morphism1 (form b) (form b).
 intros; constructor 1;
  [ apply (λx.□_b (Ext⎽b x));
- | do 2 unfold uncurry_arrows; intros; apply  (†(†H));]
+ | do 2 unfold FunClass_1_OF_Type_OF_setoid21; intros; apply  (†(†e));]
 qed.
 
+(*
 lemma xxx : ∀x.carr x → carr1 (setoid1_of_setoid x). intros; assumption; qed.
 coercion xxx nocomposites.
+*)
 
-lemma down_p : ∀S,I:SET.∀u:S⇒S.∀c:arrows1 SET I S.∀a:I.∀a':I.a=a'→u (c a)=u (c a').
-intros; unfold uncurry_arrows; change in c with (I ⇒ S);
-apply (†(†H));
+lemma down_p : ∀S:SET1.∀I:SET.∀u:S⇒S.∀c:arrows2 SET1 I S.∀a:I.∀a':I.a=a'→u (c a)=u (c a').
+intros; apply (†(†e));
 qed.
 
-alias symbol "eq" = "setoid eq".
-alias symbol "and" = "o-algebra binary meet".
-record concrete_space : Type ≝
+record concrete_space : Type2 ≝
  { bp:> BP;
    (*distr : is_distributive (form bp);*)
-   downarrow: unary_morphism (oa_P (form bp)) (oa_P (form bp));
+   downarrow: unary_morphism1 (form bp) (form bp);
    downarrow_is_sat: is_saturation ? downarrow;
    converges: ∀q1,q2.
      (Ext⎽bp q1 ∧ (Ext⎽bp q2)) = (Ext⎽bp ((downarrow q1) ∧ (downarrow q2)));
    all_covered: Ext⎽bp (oa_one (form bp)) = oa_one (concr bp);
-   il2: ∀I:SET.∀p:arrows1 SET I (oa_P (form bp)).
+   il2: ∀I:SET.∀p:arrows2 SET1 I (form bp).
      downarrow (∨ { x ∈ I | downarrow (p x) | down_p ???? }) =
      ∨ { x ∈ I | downarrow (p x) | down_p ???? };
    il1: ∀q.downarrow (A ? q) = A ? q