alias symbol "eq" = "setoid1 eq".
(* Qui, per fare le cose per bene, ci serve la nozione di funtore categorico *)
-definition o_basic_topology_of_o_basic_pair: OBP → BTop.
+definition o_basic_topology_of_o_basic_pair: OBP → OBTop.
intro t;
constructor 1;
[ apply (Oform t);
definition o_continuous_relation_of_o_relation_pair:
∀BP1,BP2.arrows2 OBP BP1 BP2 →
- arrows2 BTop (o_basic_topology_of_o_basic_pair BP1) (o_basic_topology_of_o_basic_pair BP2).
+ arrows2 OBTop (o_basic_topology_of_o_basic_pair BP1) (o_basic_topology_of_o_basic_pair BP2).
intros (BP1 BP2 t);
constructor 1;
[ apply (t \sub \f);
qed.
-definition OR : carr3 (arrows3 CAT2 OBP BTop).
+definition OR : carr3 (arrows3 CAT2 OBP OBTop).
constructor 1;
[ apply o_basic_topology_of_o_basic_pair;
| intros; constructor 1;
[ apply o_continuous_relation_of_o_relation_pair;
| apply hide;
intros; whd; unfold o_continuous_relation_of_o_relation_pair; simplify;;
- change with ((a \sub \f ⎻* ∘ A (o_basic_topology_of_o_basic_pair S)) =
- (a' \sub \f ⎻*∘A (o_basic_topology_of_o_basic_pair S)));
+ change with ((a \sub \f ⎻* ∘ oA (o_basic_topology_of_o_basic_pair S)) =
+ (a' \sub \f ⎻*∘ oA (o_basic_topology_of_o_basic_pair S)));
whd in e; cases e; clear e e2 e3 e4;
change in ⊢ (? ? ? (? ? ? ? ? % ?) ?) with ((⊩\sub S)⎻* ∘ (⊩\sub S)⎻);
apply (.= (comp_assoc2 ? ???? ?? a\sub\f⎻* ));
change in ⊢ (? ? ? (? ? ? ? ? ? %) ?) with (a'\sub\f⎻* ∘ (⊩\sub S)⎻* );
apply (.= (comp_assoc2 ? ???? ?? a'\sub\f⎻* )^-1);
apply refl2;]
-| intros 2 (o a); apply rule #;
+| intros 2 (o a); apply refl1;
| intros 6; apply refl1;]
qed.
include "o-algebra.ma".
include "o-saturations.ma".
-record basic_topology: Type2 ≝
- { carrbt:> OA;
- A: carrbt ⇒ carrbt;
- J: carrbt ⇒ carrbt;
- A_is_saturation: is_o_saturation ? A;
- J_is_reduction: is_o_reduction ? J;
- compatibility: ∀U,V. (A U >< J V) = (U >< J V)
+record Obasic_topology: Type2 ≝
+ { Ocarrbt:> OA;
+ oA: Ocarrbt ⇒ Ocarrbt;
+ oJ: Ocarrbt ⇒ Ocarrbt;
+ oA_is_saturation: is_o_saturation ? oA;
+ oJ_is_reduction: is_o_reduction ? oJ;
+ Ocompatibility: ∀U,V. (oA U >< oJ V) = (U >< oJ V)
}.
-record continuous_relation (S,T: basic_topology) : Type2 ≝
- { cont_rel:> arrows2 OA S T;
+record Ocontinuous_relation (S,T: Obasic_topology) : Type2 ≝
+ { Ocont_rel:> arrows2 OA S T;
(* reduces uses eq1, saturated uses eq!!! *)
- reduced: ∀U. U = J ? U → cont_rel U = J ? (cont_rel U);
- saturated: ∀U. U = A ? U → cont_rel⎻* U = A ? (cont_rel⎻* U)
+ Oreduced: ∀U. U = oJ ? U → Ocont_rel U = oJ ? (Ocont_rel U);
+ Osaturated: ∀U. U = oA ? U → Ocont_rel⎻* U = oA ? (Ocont_rel⎻* U)
}.
-definition continuous_relation_setoid: basic_topology → basic_topology → setoid2.
+definition Ocontinuous_relation_setoid: Obasic_topology → Obasic_topology → setoid2.
intros (S T); constructor 1;
- [ apply (continuous_relation S T)
+ [ apply (Ocontinuous_relation S T)
| constructor 1;
[ alias symbol "eq" = "setoid2 eq".
alias symbol "compose" = "category2 composition".
- apply (λr,s:continuous_relation S T. (r⎻* ) ∘ (A S) = (s⎻* ∘ (A ?)));
+ apply (λr,s:Ocontinuous_relation S T. (r⎻* ) ∘ (oA S) = (s⎻* ∘ (oA ?)));
| simplify; intros; apply refl2;
| simplify; intros; apply sym2; apply e
| simplify; intros; apply trans2; [2: apply e |3: apply e1; |1: skip]]]
qed.
-definition continuous_relation_of_continuous_relation_setoid:
- ∀P,Q. continuous_relation_setoid P Q → continuous_relation P Q ≝ λP,Q,c.c.
-coercion continuous_relation_of_continuous_relation_setoid.
+definition Ocontinuous_relation_of_Ocontinuous_relation_setoid:
+ ∀P,Q. Ocontinuous_relation_setoid P Q → Ocontinuous_relation P Q ≝ λP,Q,c.c.
+coercion Ocontinuous_relation_of_Ocontinuous_relation_setoid.
(*
theorem continuous_relation_eq':
qed.
*)
-definition continuous_relation_comp:
+definition Ocontinuous_relation_comp:
∀o1,o2,o3.
- continuous_relation_setoid o1 o2 →
- continuous_relation_setoid o2 o3 →
- continuous_relation_setoid o1 o3.
+ Ocontinuous_relation_setoid o1 o2 →
+ Ocontinuous_relation_setoid o2 o3 →
+ Ocontinuous_relation_setoid o1 o3.
intros (o1 o2 o3 r s); constructor 1;
[ apply (s ∘ r);
| intros;
apply sym1;
change in match ((s ∘ r) U) with (s (r U));
- apply (.= (reduced : ?)\sup -1);
- [ apply (.= (reduced :?)); [ assumption | apply refl1 ]
+ apply (.= (Oreduced : ?)\sup -1);
+ [ apply (.= (Oreduced :?)); [ assumption | apply refl1 ]
| apply refl1]
| intros;
apply sym1;
change in match ((s ∘ r)⎻* U) with (s⎻* (r⎻* U));
- apply (.= (saturated : ?)\sup -1);
- [ apply (.= (saturated : ?)); [ assumption | apply refl1 ]
+ apply (.= (Osaturated : ?)\sup -1);
+ [ apply (.= (Osaturated : ?)); [ assumption | apply refl1 ]
| apply refl1]]
qed.
-definition BTop: category2.
+definition OBTop: category2.
constructor 1;
- [ apply basic_topology
- | apply continuous_relation_setoid
+ [ apply Obasic_topology
+ | apply Ocontinuous_relation_setoid
| intro; constructor 1;
[ apply id2
| intros; apply e;
| intros; apply e;]
| intros; constructor 1;
- [ apply continuous_relation_comp;
+ [ apply Ocontinuous_relation_comp;
| intros; simplify;
- change with ((b⎻* ∘ a⎻* ) ∘ A o1 = ((b'⎻* ∘ a'⎻* ) ∘ A o1));
- change with (b⎻* ∘ (a⎻* ∘ A o1) = b'⎻* ∘ (a'⎻* ∘ A o1));
- change in e with (a⎻* ∘ A o1 = a'⎻* ∘ A o1);
- change in e1 with (b⎻* ∘ A o2 = b'⎻* ∘ A o2);
+ change with ((b⎻* ∘ a⎻* ) ∘ oA o1 = ((b'⎻* ∘ a'⎻* ) ∘ oA o1));
+ change with (b⎻* ∘ (a⎻* ∘ oA o1) = b'⎻* ∘ (a'⎻* ∘ oA o1));
+ change in e with (a⎻* ∘ oA o1 = a'⎻* ∘ oA o1);
+ change in e1 with (b⎻* ∘ oA o2 = b'⎻* ∘ oA o2);
apply (.= e‡#);
intro x;
- change with (eq1 ? (b⎻* (a'⎻* (A o1 x))) (b'⎻*(a'⎻* (A o1 x))));
- apply (.= †(saturated o1 o2 a' (A o1 x) ?)); [
- apply ((o_saturation_idempotent ?? (A_is_saturation o1) x)^-1);]
- apply (.= (e1 (a'⎻* (A o1 x))));
- change with (eq1 ? (b'⎻* (A o2 (a'⎻* (A o1 x)))) (b'⎻*(a'⎻* (A o1 x))));
- apply (.= †(saturated o1 o2 a' (A o1 x):?)^-1); [
- apply ((o_saturation_idempotent ?? (A_is_saturation o1) x)^-1);]
+ change with (eq1 ? (b⎻* (a'⎻* (oA o1 x))) (b'⎻*(a'⎻* (oA o1 x))));
+ apply (.= †(Osaturated o1 o2 a' (oA o1 x) ?)); [
+ apply ((o_saturation_idempotent ?? (oA_is_saturation o1) x)^-1);]
+ apply (.= (e1 (a'⎻* (oA o1 x))));
+ change with (eq1 ? (b'⎻* (oA o2 (a'⎻* (oA o1 x)))) (b'⎻*(a'⎻* (oA o1 x))));
+ apply (.= †(Osaturated o1 o2 a' (oA o1 x):?)^-1); [
+ apply ((o_saturation_idempotent ?? (oA_is_saturation o1) x)^-1);]
apply rule #;]
| intros; simplify;
- change with (((a34⎻* ∘ a23⎻* ) ∘ a12⎻* ) ∘ A o1 = ((a34⎻* ∘ (a23⎻* ∘ a12⎻* )) ∘ A o1));
+ change with (((a34⎻* ∘ a23⎻* ) ∘ a12⎻* ) ∘ oA o1 = ((a34⎻* ∘ (a23⎻* ∘ a12⎻* )) ∘ oA o1));
apply rule (#‡ASSOC ^ -1);
| intros; simplify;
- change with ((a⎻* ∘ (id2 ? o1)⎻* ) ∘ A o1 = a⎻* ∘ A o1);
+ change with ((a⎻* ∘ (id2 ? o1)⎻* ) ∘ oA o1 = a⎻* ∘ oA o1);
apply (#‡(id_neutral_right2 : ?));
| intros; simplify;
- change with (((id2 ? o2)⎻* ∘ a⎻* ) ∘ A o1 = a⎻* ∘ A o1);
+ change with (((id2 ? o2)⎻* ∘ a⎻* ) ∘ oA o1 = a⎻* ∘ oA o1);
apply (#‡(id_neutral_left2 : ?));]
qed.
-definition basic_topology_of_BTop: objs2 BTop → basic_topology ≝ λx.x.
-coercion basic_topology_of_BTop.
+definition Obasic_topology_of_OBTop: objs2 OBTop → Obasic_topology ≝ λx.x.
+coercion Obasic_topology_of_OBTop.
-definition continuous_relation_setoid_of_arrows2_BTop :
- ∀P,Q. arrows2 BTop P Q → continuous_relation_setoid P Q ≝ λP,Q,x.x.
-coercion continuous_relation_setoid_of_arrows2_BTop.
+definition Ocontinuous_relation_setoid_of_arrows2_OBTop :
+ ∀P,Q. arrows2 OBTop P Q → Ocontinuous_relation_setoid P Q ≝ λP,Q,x.x.
+coercion Ocontinuous_relation_setoid_of_arrows2_OBTop.
(*
(*CSC: unused! *)