]> matita.cs.unibo.it Git - helm.git/commitdiff
Some renaming to avoid confusion between saturations and o_saturations.
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Tue, 6 Jan 2009 15:16:49 +0000 (15:16 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Tue, 6 Jan 2009 15:16:49 +0000 (15:16 +0000)
helm/software/matita/contribs/formal_topology/overlap/o-basic_topologies.ma
helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma
helm/software/matita/contribs/formal_topology/overlap/o-saturations.ma

index 1081f5beab493cd652ed0872749ffc8f6a9f1cc3..873a9df60988a632084b8c6f24e9c2b00b619590 100644 (file)
@@ -19,8 +19,8 @@ record basic_topology: Type2 ≝
  { carrbt:> OA;
    A: carrbt ⇒ carrbt;
    J: carrbt ⇒ carrbt;
-   A_is_saturation: is_saturation ? A;
-   J_is_reduction: is_reduction ? J;
+   A_is_saturation: is_o_saturation ? A;
+   J_is_reduction: is_o_reduction ? J;
    compatibility: ∀U,V. (A U >< J V) = (U >< J V)
  }.
 
@@ -157,13 +157,13 @@ definition BTop: category2.
        alias symbol "invert" = "setoid1 symmetry".
        lapply (.= †(saturated o1 o2 a' (A o1 x) : ?));
         [3: apply (b⎻* ); | 5: apply Hletin; |1,2: skip;
-        |apply ((saturation_idempotent ?? (A_is_saturation o1) x)^-1); ]
+        |apply ((o_saturation_idempotent ?? (A_is_saturation o1) x)^-1); ]
        change in e1 with (∀x.b⎻* (A o2 x) = b'⎻* (A o2 x));
        apply (.= (e1 (a'⎻* (A o1 x))));
        alias symbol "invert" = "setoid1 symmetry".
        lapply (†((saturated ?? a' (A o1 x) : ?) ^ -1));
         [2: apply (b'⎻* ); |4: apply Hletin; | skip;
-        |apply ((saturation_idempotent ?? (A_is_saturation o1) x)^-1);]]
+        |apply ((o_saturation_idempotent ?? (A_is_saturation o1) x)^-1);]]
   | intros; simplify;
     change with (((a34⎻* ∘ a23⎻* ) ∘ a12⎻* ) ∘ A o1 = ((a34⎻* ∘ (a23⎻* ∘ a12⎻* )) ∘ A o1));
     apply rule (#‡ASSOC ^ -1);
index 6a9d7794db701429f657746dc1b41dc4922aad58..e90ec7fe401942dc219b7d98be8669760cd9fabd 100644 (file)
@@ -45,7 +45,7 @@ record concrete_space : Type2 ≝
  { bp:> BP;
    (*distr : is_distributive (form bp);*)
    downarrow: unary_morphism1 (form bp) (form bp);
-   downarrow_is_sat: is_saturation ? downarrow;
+   downarrow_is_sat: is_o_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);
index bd4e187777fb9635dc3349b9a32a6ffb2596deec..a1c83e70969009a9aabbd3338cdc6a612a30837f 100644 (file)
 include "o-algebra.ma".
 
 alias symbol "eq" = "setoid1 eq".
-definition is_saturation: ∀C:OA. unary_morphism1 C C → CProp1 ≝
+definition is_o_saturation: ∀C:OA. unary_morphism1 C C → CProp1 ≝
  λC:OA.λA:unary_morphism1 C C.
   ∀U,V. (U ≤ A V) = (A U ≤ A V).
 
-definition is_reduction: ∀C:OA. unary_morphism1 C C → CProp1 ≝
+definition is_o_reduction: ∀C:OA. unary_morphism1 C C → CProp1 ≝
  λC:OA.λJ:unary_morphism1 C C.
     ∀U,V. (J U ≤ V) = (J U ≤ J V).
 
-theorem saturation_expansive: ∀C,A. is_saturation C A → ∀U. U ≤ A U.
+theorem o_saturation_expansive: ∀C,A. is_o_saturation C A → ∀U. U ≤ A U.
  intros; apply (fi ?? (i ??)); apply (oa_leq_refl C).
 qed.
 
-theorem saturation_monotone:
- ∀C,A. is_saturation C A →
+theorem o_saturation_monotone:
+ ∀C,A. is_o_saturation C A →
   ∀U,V. U ≤ V → A U ≤ A V.
  intros; apply (if ?? (i ??)); apply (oa_leq_trans C);
-  [apply V|3: apply saturation_expansive ]
+  [apply V|3: apply o_saturation_expansive ]
  assumption.
 qed.
 
-theorem saturation_idempotent: ∀C,A. is_saturation C A → ∀U. 
+theorem o_saturation_idempotent: ∀C,A. is_o_saturation C A → ∀U. 
  eq1 C (A (A U)) (A U).
  intros; apply (oa_leq_antisym C);
   [ apply (if ?? (i (A U) U)); apply (oa_leq_refl C).
-  | apply saturation_expansive; assumption]
-qed.
\ No newline at end of file
+  | apply o_saturation_expansive; assumption]
+qed.