]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma
go notation go!
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-concrete_spaces.ma
index 7292b4ecb336c0a86580a2e5f6559bf8bc5a771f..4e989cb14f6489c02bb5d9330354e33292ad4b55 100644 (file)
 include "o-basic_pairs.ma".
 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 = (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 = (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 = (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 = (or_f_minus _ _ (rel x)).
+
+definition A : ∀b:BP. unary_morphism (oa_P (form b)) (oa_P (form b)).
+intros; constructor 1; [ apply (λx.□_b (Ext⎽b x)); | intros; apply (†(†H));] qed. 
+
 lemma xxx : ∀x.carr x → carr1 (setoid1_of_setoid x). intros; assumption; qed.
 coercion xxx.
 
+definition d_p_i : 
+  ∀S:setoid.∀I:setoid.∀d:unary_morphism S S.∀p:ums I S.ums I S.
+intros; constructor 1; [ apply (λi:I. u (c i));| intros; apply (†(†H));].
+qed. 
+
+alias symbol "eq" = "setoid eq".
+alias symbol "and" = "o-algebra binary meet".
 record concrete_space : Type ≝
  { bp:> BP;
+   (*distr : is_distributive (form bp);*)
    downarrow: unary_morphism (oa_P (form bp)) (oa_P (form bp));
    downarrow_is_sat: is_saturation ? downarrow;
    converges: ∀q1,q2.
-     or_f_minus ?? (⊩) q1 ∧ or_f_minus ?? (⊩) q2 = 
-     or_f_minus ?? (⊩) ((downarrow q1) ∧ (downarrow q2));
-   all_covered: (*⨍^-_bp*) or_f_minus ?? (⊩) (oa_one (form bp)) = oa_one (concr bp);
+     (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:setoid.∀p:ums I (oa_P (form bp)).
-     downarrow (oa_join ? I (mk_unary_morphism ?? (λi:I. downarrow (p i)) ?)) =
-     oa_join ? I (mk_unary_morphism ?? (λi:I. downarrow (p i)) ?)
+     downarrow (oa_join ? I (d_p_i ?? downarrow p)) =
+     oa_join ? I (d_p_i ?? downarrow p);
+   il1: ∀q.downarrow (A ? q) = A ? q
  }.
 
-definition bp': concrete_space → basic_pair ≝ λc.bp c.
+interpretation "o-concrete space downarrow" 'downarrow x = (fun_1 __ (downarrow _) x).
 
+definition bp': concrete_space → basic_pair ≝ λc.bp c.
 coercion bp'.
 
 record convergent_relation_pair (CS1,CS2: concrete_space) : Type ≝