definition basic_topology_of_basic_pair: basic_pair → basic_topology.
intro bp;
- letin obp ≝ (o_basic_pair_of_basic_pair bp);
- letin obt ≝ (o_basic_topology_of_o_basic_pair obp);
+ letin obt ≝ (OR (BP_to_OBP bp));
constructor 1;
[ apply (form bp);
| apply (oA obt);
∀BP1,BP2.relation_pair BP1 BP2 →
continuous_relation (basic_topology_of_basic_pair BP1) (basic_topology_of_basic_pair BP2).
intros (BP1 BP2 rp);
- letin orp ≝ (o_relation_pair_of_relation_pair ?? rp);
- letin ocr ≝ (o_continuous_relation_of_o_relation_pair ?? orp);
+ letin ocr ≝ (OR⎽⇒ (BP_to_OBP⎽⇒ rp));
constructor 1;
[ apply (rp \sub \f);
| apply (Oreduced ?? ocr);
| apply o_relation_pair_of_relation_pair_morphism_respects_comp;]
qed.
-theorem BP_to_OBP_faithful:
- ∀S,T.∀f,g:arrows2 (category2_of_category1 BP) S T.
- BP_to_OBP⎽⇒ f = BP_to_OBP⎽⇒ g → f=g.
- intros; change with ( (⊩) ∘ f \sub \c = (⊩) ∘ g \sub \c);
+theorem BP_to_OBP_faithful: faithful2 ?? BP_to_OBP.
+ intros 5 (S T); change with ( (⊩) ∘ f \sub \c = (⊩) ∘ g \sub \c);
apply (POW_faithful);
apply (.= respects_comp2 ?? POW (concr S) (concr T) (form T) f \sub \c (⊩ \sub T));
apply sym2;
apply e;
qed.
-theorem BP_to_OBP_full:
- ∀S,T.∀f. exT22 ? (λg:arrows2 ? S T. BP_to_OBP⎽⇒ g = f).
- intros;
+theorem BP_to_OBP_full: full2 ?? BP_to_OBP.
+ intros 3 (S T);
cases (POW_full (concr S) (concr T) (Oconcr_rel ?? f)) (gc Hgc);
cases (POW_full (form S) (form T) (Oform_rel ?? f)) (gf Hgf);
exists[
| cases daemon (*apply (o_relation_pair_of_relation_pair_is_morphism S T)*)]
qed.
-definition BTop_to_OBTop: carr3 (arrows3 CAT2 (category2_of_category1 BTop) OBTop).
+definition BTop_to_OBTop: carr3 ((category2_of_category1 BTop) ⇒_\c3 OBTop).
constructor 1;
[ apply o_basic_topology_of_basic_topology;
| intros; apply o_continuous_relation_of_continuous_relation_morphism;
definition functor2_setoid_of_arrows3_CAT2: ∀S,T. arrows3 CAT2 S T → functor2_setoid S T ≝ λS,T,x.x.
coercion functor2_setoid_of_arrows3_CAT2.
+notation > "B ⇒_\c3 C" right associative with precedence 72 for @{'arrows3_CAT $B $C}.
+notation "B ⇒\sub (\c 3) C" right associative with precedence 72 for @{'arrows3_CAT $B $C}.
+interpretation "'arrows3_CAT" 'arrows3_CAT A B = (arrows3 CAT2 A B).
+
definition unary_morphism_setoid: setoid → setoid → setoid.
intros;
constructor 1;
prefer coercion Type_OF_setoid. (* we prefer the lower carrier projection *)
prefer coercion Type_OF_objs1.
+
+alias symbol "exists" (instance 1) = "CProp2 exists".
+definition full2 ≝
+ λA,B:CAT2.λF:carr3 (arrows3 CAT2 A B).
+ ∀o1,o2:A.∀f.∃g:arrows2 A o1 o2.F⎽⇒ g =_2 f.
+alias symbol "exists" (instance 1) = "CProp exists".
+
+definition faithful2 ≝
+ λA,B:CAT2.λF:carr3 (arrows3 CAT2 A B).
+ ∀o1,o2:A.∀f,g:arrows2 A o1 o2.F⎽⇒ f =_2 F⎽⇒ g → f =_2 g.
+
oa_zero_bot: ∀p:oa_P.𝟘 ≤ p;
oa_one_top: ∀p:oa_P.p ≤ 𝟙;
oa_overlap_preserves_meet_: ∀p,q:oa_P.p >< q →
- p >< (⋀ { x ∈ BOOL | If x then p else q(*match x with [ true ⇒ p | false ⇒ q ]*) | IF_THEN_ELSE_p oa_P p q });
+ p >< (⋀ { x ∈ BOOL | If x then p else q | IF_THEN_ELSE_p oa_P p q });
oa_join_split: ∀I:SET.∀p.∀q:I ⇒_2 oa_P.p >< (⋁ q) = (∃i:I.p >< (q i));
(*oa_base : setoid;
1) enum non e' il nome giusto perche' non e' suriettiva
∀P,Q.arrows2 OBP P Q → Orelation_pair_setoid P Q ≝ λP,Q,c.c.
coercion Orelation_pair_setoid_of_arrows2_OBP.
+notation > "B ⇒_\obp2 C" right associative with precedence 72 for @{'arrows2_OBP $B $C}.
+notation "B ⇒\sub (\obp 2) C" right associative with precedence 72 for @{'arrows2_OBP $B $C}.
+interpretation "'arrows2_OBP" 'arrows2_OBP A B = (arrows2 OBP A B).
+
(*
definition BPext: ∀o: BP. form o ⇒ Ω \sup (concr o).
intros; constructor 1;
qed.
definition o_continuous_relation_of_o_relation_pair:
- ∀BP1,BP2.arrows2 OBP BP1 BP2 →
- arrows2 OBTop (o_basic_topology_of_o_basic_pair BP1) (o_basic_topology_of_o_basic_pair BP2).
+ ∀BP1,BP2.BP1 ⇒_\obp2 BP2 →
+ (o_basic_topology_of_o_basic_pair BP1) ⇒_\obt2 (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 OBTop).
+definition OR : carr3 (OBP ⇒_\c3 OBTop).
constructor 1;
[ apply o_basic_topology_of_o_basic_pair;
| intros; constructor 1;
∀P,Q. arrows2 OBTop P Q → Ocontinuous_relation_setoid P Q ≝ λP,Q,x.x.
coercion Ocontinuous_relation_setoid_of_arrows2_OBTop.
+notation > "B ⇒_\obt2 C" right associative with precedence 72 for @{'arrows2_OBT $B $C}.
+notation "B ⇒\sub (\obt 2) C" right associative with precedence 72 for @{'arrows2_OBT $B $C}.
+interpretation "'arrows2_OBT" 'arrows2_OBT A B = (arrows2 OBTop A B).
+
+
(*
(*CSC: unused! *)
(* this proof is more logic-oriented than set/lattice oriented *)
notation > "B ⇒_\r1 C" right associative with precedence 72 for @{'arrows1_REL $B $C}.
notation "B ⇒\sub (\r 1) C" right associative with precedence 72 for @{'arrows1_REL $B $C}.
-interpretation "'arrows1_SET" 'arrows1_REL A B = (arrows1 REL A B).
+interpretation "'arrows1_REL" 'arrows1_REL A B = (arrows1 REL A B).
+notation > "B ⇒_\r2 C" right associative with precedence 72 for @{'arrows2_REL $B $C}.
+notation "B ⇒\sub (\r 2) C" right associative with precedence 72 for @{'arrows2_REL $B $C}.
+interpretation "'arrows2_REL" 'arrows2_REL A B = (arrows2 (category2_of_category1 REL) A B).
definition full_subset: ∀s:REL. Ω^s.
| apply orelation_of_relation_preserves_composition; ]
qed.
-theorem POW_faithful:
- ∀S,T.∀f,g:arrows2 (category2_of_category1 REL) S T.
- POW⎽⇒ f =_2 POW⎽⇒ g → f =_2 g.
- intros; unfold POW in e; simplify in e; cases e;
+theorem POW_faithful: faithful2 ?? POW.
+ intros 5; unfold POW in e; simplify in e; cases e;
unfold orelation_of_relation in e3; simplify in e3; clear e e1 e2 e4;
intros 2; cases (e3 {(x)});
split; intro; [ lapply (s y); | lapply (s1 y); ]
|*: cases Hletin; cases x1; change in f3 with (x =_1 w); apply (. f3‡#); assumption;]
qed.
-
+(*
lemma currify: ∀A,B,C. (A × B ⇒_1 C) → A → (B ⇒_1 C).
intros; constructor 1; [ apply (b c); | intros; apply (#‡e);]
qed.
+*)
-theorem POW_full: ∀S,T.∀f: (POW S) ⇒_\o2 (POW T) . exT22 ? (λg. POW⎽⇒ g = f).
- intros; exists;
+theorem POW_full: full2 ?? POW.
+ intros 3 (S T); exists;
[ constructor 1; constructor 1;
[ apply (λx:carr S.λy:carr T. y ∈ f {(x)});
| intros; unfold FunClass_1_OF_carr2; lapply (.= e1‡#);