pc1/props pc1_t
pc1/props pc1_pr0_u2
pc1/props pc1_head
+
pc3/dec pc3_dec
pc3/dec pc3_abst_dec
pc3/fsubst0 pc3_pr2_fsubst0
pc3/nf2 pc3_nf2
pc3/nf2 pc3_nf2_unfold
pc3/pc1 pc3_pc1
-pc3/props clear_pc3_trans
pc3/props pc3_pr2_r
pc3/props pc3_pr2_x
pc3/props pc3_pr3_r
pc3/props pc3_pr3_x
pc3/props pc3_pr3_t
-pc3/props pc3_refl
pc3/props pc3_s
pc3/props pc3_thin_dx
pc3/props pc3_head_1
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "Basic_2/reducibility/cpr.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL CONVERSION ON TERMS ***************************)
+
+definition cpc: lenv → relation term ≝
+ λL,T1,T2. L ⊢ T1 ➡ T2 ∨ L ⊢ T2 ➡ T1.
+
+interpretation
+ "context-sensitive parallel conversion (term)"
+ 'PConv L T1 T2 = (cpc L T1 T2).
+
+(* Basic properties *********************************************************)
+
+lemma cpc_refl: ∀L,T. L ⊢ T ⬌ T.
+/2 width=1/ qed.
+
+lemma cpc_sym: ∀L,T1,T2. L ⊢ T1 ⬌ T2 → L ⊢ T2 ⬌ T1.
+#L #T1 #T2 * /2 width=1/
+qed.
+
+lemma cpc_cpr: ∀L,T1,T2. L ⊢ T1 ⬌ T2 → ∃∃T. L ⊢ T1 ➡ T & L ⊢ T2 ➡ T.
+#L #T1 #T2 * /2 width=3/
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "Basic_2/conversion/cpc.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL CONVERSION ON TERMS ***************************)
+
+(* Main properties **********************************************************)
+
+theorem cpc_conf: ∀L,T0,T1,T2. L ⊢ T0 ⬌ T1 → L ⊢ T0 ⬌ T2 →
+ ∃∃T. L ⊢ T1 ⬌ T & L ⊢ T2 ⬌ T.
+/3 width=3/ qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "Basic_2/conversion/cpc.ma".
+
+(* CONTEXT-SENSITIVE PARALLEL EQUIVALENCE ON TERMS **************************)
+
+definition cpcs: lenv → relation term ≝
+ λL. TC … (cpc L).
+
+interpretation "context-sensitive parallel equivalence (term)"
+ 'PConvStar L T1 T2 = (cpcs L T1 T2).
+
+(* Basic eliminators ********************************************************)
+
+lemma cpcs_ind: ∀L,T1. ∀R:predicate term. R T1 →
+ (∀T,T2. L ⊢ T1 ⬌* T → L ⊢ T ⬌ T2 → R T → R T2) →
+ ∀T2. L ⊢ T1 ⬌* T2 → R T2.
+#L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) //
+qed-.
+
+(* Basic properties *********************************************************)
+
+(* Basic_1: was: pc3_refl *)
+lemma cpcs_refl: ∀L,T. L ⊢ T ⬌* T.
+/2 width=1/ qed.
+
+lemma cpcs_strap1: ∀L,T1,T,T2.
+ L ⊢ T1 ⬌* T → L ⊢ T ⬌ T2 → L ⊢ T1 ⬌* T2.
+/2 width=3/ qed.
+
+lemma cpcs_strap2: ∀L,T1,T,T2.
+ L ⊢ T1 ⬌ T → L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2.
+/2 width=3/ qed.
+
+(* Basic_1: removed theorems 1: clear_pc3_trans *)
%.stats: CHARS = $(shell cat $(MAS) | wc -c)
-%.stats: C1 = $(shell grep "inductive \|record " $(MAS) | wc -l)
-%.stats: C2 = $(shell grep "definition \|let rec " $(MAS) | wc -l)
-%.stats: C3 = $(shell grep "inductive \|record \|definition \|let rec " $(MAS) | wc -l)
-%.stats: P1 = $(shell grep "theorem " $(MAS) | wc -l)
-%.stats: P2 = $(shell grep "lemma " $(MAS) | wc -l)
-%.stats: P3 = $(shell grep "lemma \|theorem " $(MAS) | wc -l)
-
-%.stats: TBL = ld_$*_sum
-
%.stats:
@printf '\x1B[1;40;37m'
@printf '%-15s %-42s' 'Statistics for:' $*
@printf ' %-6s %3i' Marks `grep "(\*\*)" $(MAS) | wc -l`
@printf ' %-10s' ''
@printf '\x1B[0m\n'
- @printf 'name "$(TBL)"\n\n' > $*/$(TBL).tbl
- @printf 'table {\n' >> $*/$(TBL).tbl
- @printf ' class "grey" [ "category"\n' >> $*/$(TBL).tbl
- @printf ' [ "objects" * ]\n' >> $*/$(TBL).tbl
- @printf ' ]\n' >> $*/$(TBL).tbl
- @printf ' class "green" [ "propositions"\n' >> $*/$(TBL).tbl
- @printf ' [ "theorems" "$(P1)" * ]\n' >> $*/$(TBL).tbl
- @printf ' [ "lemmas" "$(P2)" * ]\n' >> $*/$(TBL).tbl
- @printf ' [ "total" "$(P3)" * ]\n' >> $*/$(TBL).tbl
- @printf ' ]\n' >> $*/$(TBL).tbl
- @printf ' class "yellow" [ "concepts"\n' >> $*/$(TBL).tbl
- @printf ' [ "declared" "$(C1)" * ]\n' >> $*/$(TBL).tbl
- @printf ' [ "defined" "$(C2)" * ]\n' >> $*/$(TBL).tbl
- @printf ' [ "total" "$(C3)" * ]\n' >> $*/$(TBL).tbl
- @printf ' ]\n' >> $*/$(TBL).tbl
- @printf '}\n\n' >> $*/$(TBL).tbl
- @printf 'class "component" { 0 }\n\n' >> $*/$(TBL).tbl
- @printf 'class "plane" { 1 } { 3 } { 5 } \n\n' >> $*/$(TBL).tbl
- @printf 'class "number" { 2 } { 4 } { 6 } \n\n' >> $*/$(TBL).tbl
+
+# sum ########################################################################
+
+tbls: $(PACKAGES:%=etc/ld_%_sum.tbl)
+
+etc/ld_%_sum.tbl: MAS = $(shell find $* -name "*.ma")
+
+%.tbl: C1 = $(shell grep "inductive \|record " $(MAS) | wc -l)
+%.tbl: C2 = $(shell grep "definition \|let rec " $(MAS) | wc -l)
+%.tbl: C3 = $(shell grep "inductive \|record \|definition \|let rec " $(MAS) | wc -l)
+%.tbl: P1 = $(shell grep "theorem " $(MAS) | wc -l)
+%.tbl: P2 = $(shell grep "lemma " $(MAS) | wc -l)
+%.tbl: P3 = $(shell grep "lemma \|theorem " $(MAS) | wc -l)
+
+etc/ld_%_sum.tbl: $(MAS)
+ @mkdir -p etc
+ @printf 'Summary for: $*\n'
+ @printf 'name "$(basename $(@F))"\n\n' > $@
+ @printf 'table {\n' >> $@
+ @printf ' class "grey" [ "category"\n' >> $@
+ @printf ' [ "objects" * ]\n' >> $@
+ @printf ' ]\n' >> $@
+ @printf ' class "green" [ "propositions"\n' >> $@
+ @printf ' [ "theorems" "$(P1)" * ]\n' >> $@
+ @printf ' [ "lemmas" "$(P2)" * ]\n' >> $@
+ @printf ' [ "total" "$(P3)" * ]\n' >> $@
+ @printf ' ]\n' >> $@
+ @printf ' class "yellow" [ "concepts"\n' >> $@
+ @printf ' [ "declared" "$(C1)" * ]\n' >> $@
+ @printf ' [ "defined" "$(C2)" * ]\n' >> $@
+ @printf ' [ "total" "$(C3)" * ]\n' >> $@
+ @printf ' ]\n' >> $@
+ @printf '}\n\n' >> $@
+ @printf 'class "component" { 0 }\n\n' >> $@
+ @printf 'class "plane" { 1 } { 3 } { 5 }\n\n' >> $@
+ @printf 'class "number" { 2 } { 4 } { 6 }\n\n' >> $@
universe constraint Type[2] < Type[3].
universe constraint Type[3] < Type[4].
universe constraint Type[4] < Type[5].
+universe constraint Type[5] < Type[6].
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/categories.ma".
+include "formal_topology/notation.ma".
+
+record Fo (C1,C2:CAT2) (F:arrows3 CAT2 C1 C2) : Type[2] ≝ {
+ F2: C2;
+ F1: C1;
+ FP: map_objs2 ?? F F1 =_\ID F2
+}.
+
+notation "ℱ\sub 1 x" non associative with precedence 60 for @{'F1 $x}.
+notation > "ℱ_1" non associative with precedence 90 for @{F1 ???}.
+interpretation "F1" 'F1 x = (F1 ??? x).
+
+notation "ℱ\sub 2 x" non associative with precedence 60 for @{'F2 $x}.
+notation > "ℱ_2" non associative with precedence 90 for @{F2 ???}.
+interpretation "F2" 'F2 x = (F2 ??? x).
+
+lemma REW : ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.∀X,Y:Fo ?? F.
+ arrows2 C2 (F (ℱ_1 X)) (F (ℱ_1 Y)) →
+ arrows2 C2 (ℱ_2 X) (ℱ_2 Y).
+intros 5; cases X; cases Y; clear X Y;
+cases H; cases H1; intros; assumption;
+qed.
+
+record Fm_c (C1,C2:CAT2) (F:arrows3 CAT2 C1 C2) (X,Y:Fo ?? F) : Type[2] ≝ {
+ Fm2: arrows2 C2 (F2 ??? X) (F2 ??? Y);
+ Fm1: arrows2 C1 (F1 ??? X) (F1 ??? Y);
+ FmP: REW ?? F X Y (map_arrows2 ?? F ?? Fm1) = Fm2
+}.
+
+notation "ℳ\sub 1 x" non associative with precedence 60 for @{'Fm1 $x}.
+notation > "ℳ_1" non associative with precedence 90 for @{Fm1 ?????}.
+interpretation "Fm1" 'Fm1 x = (Fm1 ????? x).
+
+notation "ℳ\sub 2 x" non associative with precedence 60 for @{'Fm2 $x}.
+notation > "ℳ_2" non associative with precedence 90 for @{Fm2 ?????}.
+interpretation "Fm2" 'Fm2 x = (Fm2 ????? x).
+
+definition Fm :
+ ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.
+ Fo ?? F → Fo ?? F → setoid2.
+intros (C1 C2 F X Y); constructor 1; [apply (Fm_c C1 C2 F X Y)]
+constructor 1; [apply (λf,g.Fm2 ????? f =_2 Fm2 ????? g);]
+[ intro; apply refl2;
+| intros 3; apply sym2; assumption;
+| intros 5; apply (trans2 ?? ??? x1 x2);]
+qed.
+
+definition F_id :
+ ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.∀o.Fm ?? F o o.
+intros; constructor 1;
+ [ apply (id2 C2 (F2 ??? o));
+ | apply (id2 C1 (F1 ??? o));
+ | cases o; cases H; simplify; apply (respects_id2 ?? F);]
+qed.
+
+definition F_comp :
+ ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.∀o1,o2,o3.
+ (Fm ?? F o1 o2) × (Fm ?? F o2 o3) ⇒_2 (Fm ?? F o1 o3).
+intros; constructor 1;
+[ intros (f g); constructor 1;
+ [ apply (comp2 C2 ??? (ℳ_2 f) (ℳ_2 g));
+ | apply (comp2 C1 ??? (ℳ_1 f) (ℳ_1 g));
+ | apply hide; cases o1 in f; cases o2 in g; cases o3; clear o1 o2 o3;
+ cases H; cases H1; cases H2; intros 2; cases c; cases c1; clear c c1;
+ simplify; apply (.= (respects_comp2:?)); apply (e1‡e);]
+| intros 6; change with ((ℳ_2 b ∘ ℳ_2 a) = (ℳ_2 b' ∘ ℳ_2 a'));
+ change in e1 with (ℳ_2 b = ℳ_2 b');
+ change in e with (ℳ_2 a = ℳ_2 a');
+ apply (e‡e1);]
+qed.
+
+
+definition Apply : ∀C1,C2: CAT2.arrows3 CAT2 C1 C2 → CAT2.
+intros (C1 C2 F);
+constructor 1;
+[ apply (Fo ?? F);
+| apply (Fm ?? F);
+| apply F_id;
+| apply F_comp;
+| intros; apply (comp_assoc2 C2 ???? (ℳ_2 a12) (ℳ_2 a23) (ℳ_2 a34));
+| intros; apply (id_neutral_right2 C2 ?? (ℳ_2 a));
+| intros; apply (id_neutral_left2 C2 ?? (ℳ_2 a));]
+qed.
+
+definition faithful ≝
+ λC1,C2.λF:arrows3 CAT2 C1 C2.∀S,T.∀f,g:arrows2 C1 S T.
+ map_arrows2 ?? F ?? f = map_arrows2 ?? F ?? g → f=g.
+
+definition Ylppa : ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.
+ faithful ?? F → let rC2 ≝ Apply ?? F in arrows3 CAT2 rC2 C1.
+intros; constructor 1;
+[ intro; apply (ℱ_1 o);
+| intros; constructor 1;
+ [ intros; apply (ℳ_1 c);
+ | apply hide; intros; apply f; change in e with (ℳ_2 a = ℳ_2 a');
+ lapply (FmP ????? a) as H1; lapply (FmP ????? a') as H2;
+ cut (REW ????? (map_arrows2 ?? F ?? (ℳ_1 a)) =
+ REW ????? (map_arrows2 ?? F ?? (ℳ_1 a')));[2:
+ apply (.= H1); apply (.= e); apply (H2^-1);]
+ clear H1 H2 e; cases S in a a' Hcut; cases T;
+ cases H; cases H1; simplify; intros; assumption;]
+| intro; apply rule #;
+| intros; simplify; apply rule #;]
+qed.
+
+
+
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/relations.ma".
+include "formal_topology/notation.ma".
+
+record basic_pair: Type[1] ≝ {
+ concr: REL; form: REL; rel: concr ⇒_\r1 form
+}.
+
+interpretation "basic pair relation" 'Vdash2 x y c = (fun21 ??? (rel c) x y).
+interpretation "basic pair relation (non applied)" 'Vdash c = (rel c).
+
+record relation_pair (BP1,BP2: basic_pair): Type[1] ≝ {
+ concr_rel: (concr BP1) ⇒_\r1 (concr BP2); form_rel: (form BP1) ⇒_\r1 (form BP2);
+ commute: comp1 REL ??? concr_rel (rel ?) =_1 form_rel ∘ ⊩
+ }.
+
+interpretation "concrete relation" 'concr_rel r = (concr_rel ?? r).
+interpretation "formal relation" 'form_rel r = (form_rel ?? r).
+
+definition relation_pair_equality: ∀o1,o2. equivalence_relation1 (relation_pair o1 o2).
+ intros; constructor 1; [ apply (λr,r'. ⊩ ∘ r \sub\c = ⊩ ∘ r' \sub\c);
+ | simplify; intros; apply refl1;
+ | simplify; intros 2; apply sym1;
+ | simplify; intros 3; apply trans1; ]
+qed.
+
+definition relation_pair_setoid: basic_pair → basic_pair → setoid1.
+ intros;
+ constructor 1;
+ [ apply (relation_pair b b1)
+ | apply relation_pair_equality
+ ]
+qed.
+
+definition relation_pair_of_relation_pair_setoid :
+ ∀P,Q. relation_pair_setoid P Q → relation_pair P Q ≝ λP,Q,x.x.
+coercion relation_pair_of_relation_pair_setoid.
+
+alias symbol "compose" (instance 1) = "category1 composition".
+lemma eq_to_eq':
+ ∀o1,o2.∀r,r':relation_pair_setoid o1 o2. r =_1 r' → r \sub\f ∘ ⊩ =_1 r'\sub\f ∘ ⊩.
+ intros 5 (o1 o2 r r' H);
+ apply (.= (commute ?? r)^-1);
+ change in H with (⊩ ∘ r \sub\c = ⊩ ∘ r' \sub\c);
+ apply rule (.= H);
+ apply (commute ?? r').
+qed.
+
+definition id_relation_pair: ∀o:basic_pair. relation_pair o o.
+ intro;
+ constructor 1;
+ [1,2: apply id1;
+ | lapply (id_neutral_right1 ? (concr o) ? (⊩)) as H;
+ lapply (id_neutral_left1 ?? (form o) (⊩)) as H1;
+ apply (.= H);
+ apply (H1^-1);]
+qed.
+
+lemma relation_pair_composition:
+ ∀o1,o2,o3: basic_pair.
+ relation_pair_setoid o1 o2 → relation_pair_setoid o2 o3 → relation_pair_setoid o1 o3.
+intros 3 (o1 o2 o3);
+ intros (r r1);
+ constructor 1;
+ [ apply (r1 \sub\c ∘ r \sub\c)
+ | apply (r1 \sub\f ∘ r \sub\f)
+ | lapply (commute ?? r) as H;
+ lapply (commute ?? r1) as H1;
+ alias symbol "trans" = "trans1".
+ alias symbol "assoc" = "category1 assoc".
+ apply (.= ASSOC);
+ apply (.= #‡H1);
+ alias symbol "invert" = "setoid1 symmetry".
+ apply (.= ASSOC ^ -1);
+ apply (.= H‡#);
+ apply ASSOC]
+qed.
+
+lemma relation_pair_composition_is_morphism:
+ ∀o1,o2,o3: basic_pair.
+ ∀a,a':relation_pair_setoid o1 o2.
+ ∀b,b':relation_pair_setoid o2 o3.
+ a=a' → b=b' →
+ relation_pair_composition o1 o2 o3 a b
+ = relation_pair_composition o1 o2 o3 a' b'.
+intros 3 (o1 o2 o3);
+ intros;
+ change with (⊩ ∘ (b\sub\c ∘ a\sub\c) = ⊩ ∘ (b'\sub\c ∘ a'\sub\c));
+ change in e with (⊩ ∘ a \sub\c = ⊩ ∘ a' \sub\c);
+ change in e1 with (⊩ ∘ b \sub\c = ⊩ ∘ b' \sub\c);
+ apply (.= ASSOC);
+ apply (.= #‡e1);
+ apply (.= #‡(commute ?? b'));
+ apply (.= ASSOC ^ -1);
+ apply (.= e‡#);
+ apply (.= ASSOC);
+ apply (.= #‡(commute ?? b')^-1);
+ apply (ASSOC ^ -1);
+qed.
+
+definition relation_pair_composition_morphism:
+ ∀o1,o2,o3. binary_morphism1 (relation_pair_setoid o1 o2) (relation_pair_setoid o2 o3) (relation_pair_setoid o1 o3).
+ intros;
+ constructor 1;
+ [ apply relation_pair_composition;
+ | apply relation_pair_composition_is_morphism;]
+qed.
+
+lemma relation_pair_composition_morphism_assoc:
+Πo1:basic_pair
+.Πo2:basic_pair
+ .Πo3:basic_pair
+ .Πo4:basic_pair
+ .Πa12:relation_pair_setoid o1 o2
+ .Πa23:relation_pair_setoid o2 o3
+ .Πa34:relation_pair_setoid o3 o4
+ .relation_pair_composition_morphism o1 o3 o4
+ (relation_pair_composition_morphism o1 o2 o3 a12 a23) a34
+ =relation_pair_composition_morphism o1 o2 o4 a12
+ (relation_pair_composition_morphism o2 o3 o4 a23 a34).
+ intros;
+ change with (⊩ ∘ (a34\sub\c ∘ (a23\sub\c ∘ a12\sub\c)) =
+ ⊩ ∘ ((a34\sub\c ∘ a23\sub\c) ∘ a12\sub\c));
+ alias symbol "refl" = "refl1".
+ alias symbol "prop2" = "prop21".
+ apply (ASSOC‡#);
+qed.
+
+lemma relation_pair_composition_morphism_respects_id:
+ ∀o1,o2:basic_pair.∀a:relation_pair_setoid o1 o2.
+ relation_pair_composition_morphism o1 o1 o2 (id_relation_pair o1) a=a.
+ intros;
+ change with (⊩ ∘ (a\sub\c ∘ (id_relation_pair o1)\sub\c) = ⊩ ∘ a\sub\c);
+ apply ((id_neutral_right1 ????)‡#);
+qed.
+
+lemma relation_pair_composition_morphism_respects_id_r:
+ ∀o1,o2:basic_pair.∀a:relation_pair_setoid o1 o2.
+ relation_pair_composition_morphism o1 o2 o2 a (id_relation_pair o2)=a.
+ intros;
+ change with (⊩ ∘ ((id_relation_pair o2)\sub\c ∘ a\sub\c) = ⊩ ∘ a\sub\c);
+ apply ((id_neutral_left1 ????)‡#);
+qed.
+
+definition BP: category1.
+ constructor 1;
+ [ apply basic_pair
+ | apply relation_pair_setoid
+ | apply id_relation_pair
+ | apply relation_pair_composition_morphism
+ | apply relation_pair_composition_morphism_assoc;
+ | apply relation_pair_composition_morphism_respects_id;
+ | apply relation_pair_composition_morphism_respects_id_r;]
+qed.
+
+definition basic_pair_of_BP : objs1 BP → basic_pair ≝ λx.x.
+coercion basic_pair_of_BP.
+
+definition relation_pair_setoid_of_arrows1_BP :
+ ∀P,Q. arrows1 BP P Q → relation_pair_setoid P Q ≝ λP,Q,x.x.
+coercion relation_pair_setoid_of_arrows1_BP.
+
+(*
+definition BPext: ∀o: BP. (form o) ⇒_1 Ω^(concr o).
+ intros; constructor 1;
+ [ apply (ext ? ? (rel o));
+ | intros;
+ apply (.= #‡e);
+ apply refl1]
+qed.
+
+definition BPextS: ∀o: BP. Ω^(form o) ⇒_1 Ω^(concr o).
+ intros; constructor 1;
+ [ apply (minus_image ?? (rel o));
+ | intros; apply (#‡e); ]
+qed.
+
+definition fintersects: ∀o: BP. (form o) × (form o) ⇒_1 Ω^(form o).
+ intros (o); constructor 1;
+ [ apply (λa,b: form o.{c | BPext o c ⊆ BPext o a ∩ BPext o b });
+ intros; simplify; apply (.= (†e)‡#); apply refl1
+ | intros; split; simplify; intros;
+ [ apply (. #‡((†e^-1)‡(†e1^-1))); assumption
+ | apply (. #‡((†e)‡(†e1))); assumption]]
+qed.
+
+interpretation "fintersects" 'fintersects U V = (fun21 ??? (fintersects ?) U V).
+
+definition fintersectsS:
+ ∀o:BP. Ω^(form o) × Ω^(form o) ⇒_1 Ω^(form o).
+ intros (o); constructor 1;
+ [ apply (λo: basic_pair.λa,b: Ω^(form o).{c | BPext o c ⊆ BPextS o a ∩ BPextS o b });
+ intros; simplify; apply (.= (†e)‡#); apply refl1
+ | intros; split; simplify; intros;
+ [ apply (. #‡((†e^-1)‡(†e1^-1))); assumption
+ | apply (. #‡((†e)‡(†e1))); assumption]]
+qed.
+
+interpretation "fintersectsS" 'fintersects U V = (fun21 ??? (fintersectsS ?) U V).
+
+definition relS: ∀o: BP. (concr o) × Ω^(form o) ⇒_1 CPROP.
+ intros (o); constructor 1;
+ [ apply (λx:concr o.λS: Ω^(form o).∃y:form o.y ∈ S ∧ x ⊩⎽o y);
+ | intros; split; intros; cases e2; exists [1,3: apply w]
+ [ apply (. (#‡e1^-1)‡(e^-1‡#)); assumption
+ | apply (. (#‡e1)‡(e‡#)); assumption]]
+qed.
+
+interpretation "basic pair relation for subsets" 'Vdash2 x y c = (fun21 (concr ?) ?? (relS c) x y).
+interpretation "basic pair relation for subsets (non applied)" 'Vdash c = (fun21 ??? (relS c)).
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/basic_pairs_to_o-basic_pairs.ma".
+include "formal_topology/o-basic_pairs_to_o-basic_topologies.ma".
+include "formal_topology/basic_pairs.ma".
+include "formal_topology/basic_topologies.ma".
+
+definition basic_topology_of_basic_pair: basic_pair → basic_topology.
+ intro bp;
+ letin obt ≝ (OR (BP_to_OBP bp));
+ constructor 1;
+ [ apply (form bp);
+ | apply (oA obt);
+ | apply (oJ obt);
+ | apply (oA_is_saturation obt);
+ | apply (oJ_is_reduction obt);
+ | apply (Ocompatibility obt); ]
+qed.
+
+definition continuous_relation_of_relation_pair:
+ ∀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 ocr ≝ (OR⎽⇒ (BP_to_OBP⎽⇒ rp));
+ constructor 1;
+ [ apply (rp \sub \f);
+ | apply (Oreduced ?? ocr);
+ | apply (Osaturated ?? ocr); ]
+qed.
+
+alias symbol "compose" (instance 3) = "category1 composition".
+alias symbol "compose" (instance 3) = "category1 composition".
+record functor1 (C1: category1) (C2: category1) : Type[2] ≝
+ { map_objs1:1> C1 → C2;
+ map_arrows1: ∀S,T. unary_morphism1 (arrows1 ? S T) (arrows1 ? (map_objs1 S) (map_objs1 T));
+ respects_id1: ∀o:C1. map_arrows1 ?? (id1 ? o) =_1 id1 ? (map_objs1 o);
+ respects_comp1:
+ ∀o1,o2,o3.∀f1:arrows1 ? o1 o2.∀f2:arrows1 ? o2 o3.
+ map_arrows1 ?? (f2 ∘ f1) =_1 map_arrows1 ?? f2 ∘ map_arrows1 ?? f1}.
+
+(*
+definition BTop_of_BP: functor1 BP BTop.
+ constructor 1;
+ [ apply basic_topology_of_basic_pair
+ | intros; constructor 1 [ apply continuous_relation_of_relation_pair; ]
+ | simplify; intro;
+ ]
+qed.
+
+lemma BBBB_faithful : failthful2 ?? VVV
+FIXME
+*)
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/basic_pairs.ma".
+include "formal_topology/o-basic_pairs.ma".
+include "formal_topology/relations_to_o-algebra.ma".
+
+definition o_basic_pair_of_basic_pair: basic_pair → Obasic_pair.
+ intro b;
+ constructor 1;
+ [ apply (POW (concr b));
+ | apply (POW (form b));
+ | apply (POW⎽⇒ ?); apply (rel b); ]
+qed.
+
+definition o_relation_pair_of_relation_pair:
+ ∀BP1,BP2. relation_pair BP1 BP2 →
+ Orelation_pair (o_basic_pair_of_basic_pair BP1) (o_basic_pair_of_basic_pair BP2).
+ intros;
+ constructor 1;
+ [ unfold o_basic_pair_of_basic_pair; simplify; apply (POW⎽⇒ ?); apply (r\sub \c);
+ | apply (map_arrows2 ?? POW (form BP1) (form BP2) (r \sub \f));
+ | apply (.= (respects_comp2 ?? POW (concr BP1) (concr BP2) (form BP2) r\sub\c (⊩\sub BP2) )^-1);
+ cut ( ⊩ \sub BP2 ∘ r \sub \c =_12 r\sub\f ∘ ⊩ \sub BP1) as H;
+ [ apply (.= †H);
+ apply (respects_comp2 ?? POW (concr BP1) (form BP1) (form BP2) (⊩\sub BP1) r\sub\f);
+ | apply commute;]]
+qed.
+
+lemma o_relation_pair_of_relation_pair_is_morphism :
+ ∀S,T:category2_of_category1 BP.
+ ∀a,b:arrows2 (category2_of_category1 BP) S T.a=b →
+ (eq2 (arrows2 OBP (o_basic_pair_of_basic_pair S) (o_basic_pair_of_basic_pair T)))
+ (o_relation_pair_of_relation_pair S T a) (o_relation_pair_of_relation_pair S T b).
+intros 2 (S T);
+ intros (a b Eab); split; unfold o_relation_pair_of_relation_pair; simplify;
+ unfold o_basic_pair_of_basic_pair; simplify;
+ [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x);
+ | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
+ | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
+ | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
+ simplify;
+ apply (prop12);
+ apply (.= (respects_comp2 ?? POW (concr S) (concr T) (form T) (a\sub\c) (⊩\sub T))^-1);
+ apply sym2;
+ apply (.= (respects_comp2 ?? POW (concr S) (concr T) (form T) (b\sub\c) (⊩\sub T))^-1);
+ apply sym2;
+ apply prop12;
+ apply Eab;
+qed.
+
+lemma o_relation_pair_of_relation_pair_morphism :
+ ∀S,T:category2_of_category1 BP.
+ unary_morphism2 (arrows2 (category2_of_category1 BP) S T)
+ (arrows2 OBP (o_basic_pair_of_basic_pair S) (o_basic_pair_of_basic_pair T)).
+intros (S T);
+ constructor 1;
+ [ apply (o_relation_pair_of_relation_pair S T);
+ | apply (o_relation_pair_of_relation_pair_is_morphism S T)]
+qed.
+
+lemma o_relation_pair_of_relation_pair_morphism_respects_id:
+ ∀o:category2_of_category1 BP.
+ o_relation_pair_of_relation_pair_morphism o o (id2 (category2_of_category1 BP) o)
+ = id2 OBP (o_basic_pair_of_basic_pair o).
+ simplify; intros; whd; split;
+ [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x);
+ | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
+ | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
+ | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
+ simplify;
+ apply prop12;
+ apply prop22;[2,4,6,8: apply rule #;]
+ apply (respects_id2 ?? POW (concr o));
+qed.
+
+lemma o_relation_pair_of_relation_pair_morphism_respects_comp:
+ ∀o1,o2,o3:category2_of_category1 BP.
+ ∀f1:arrows2 (category2_of_category1 BP) o1 o2.
+ ∀f2:arrows2 (category2_of_category1 BP) o2 o3.
+ (eq2 (arrows2 OBP (o_basic_pair_of_basic_pair o1) (o_basic_pair_of_basic_pair o3)))
+ (o_relation_pair_of_relation_pair_morphism o1 o3 (f2 ∘ f1))
+ (comp2 OBP ???
+ (o_relation_pair_of_relation_pair_morphism o1 o2 f1)
+ (o_relation_pair_of_relation_pair_morphism o2 o3 f2)).
+ simplify; intros; whd; split;
+ [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x);
+ | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
+ | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
+ | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
+ simplify;
+ apply prop12;
+ apply prop22;[2,4,6,8: apply rule #;]
+ apply (respects_comp2 ?? POW (concr o1) (concr o2) (concr o3) f1\sub\c f2\sub\c);
+qed.
+
+definition BP_to_OBP: carr3 (arrows3 CAT2 (category2_of_category1 BP) OBP).
+ constructor 1;
+ [ apply o_basic_pair_of_basic_pair;
+ | intros; apply o_relation_pair_of_relation_pair_morphism;
+ | apply o_relation_pair_of_relation_pair_morphism_respects_id;
+ | apply o_relation_pair_of_relation_pair_morphism_respects_comp;]
+qed.
+
+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 (.= respects_comp2 ?? POW (concr S) (concr T) (form T) g \sub \c (⊩ \sub T));
+ apply sym2;
+ apply e;
+qed.
+
+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[
+ constructor 1; [apply gc|apply gf]
+ apply (POW_faithful);
+ apply (let xxxx ≝POW in .= respects_comp2 ?? POW (concr S) (concr T) (form T) gc (rel T));
+ apply rule (.= Hgc‡#);
+ apply (.= Ocommute ?? f);
+ apply (.= #‡Hgf^-1);
+ apply (let xxxx ≝POW in (respects_comp2 ?? POW (concr S) (form S) (form T) (rel S) gf)^-1)]
+ split;
+ [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x);
+ | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
+ | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
+ | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
+ simplify; apply (†(Hgc‡#));
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/relations.ma".
+include "formal_topology/saturations.ma".
+
+record basic_topology: Type[1] ≝
+ { carrbt:> REL;
+ A: Ω^carrbt ⇒_1 Ω^carrbt;
+ J: Ω^carrbt ⇒_1 Ω^carrbt;
+ A_is_saturation: is_saturation ? A;
+ J_is_reduction: is_reduction ? J;
+ compatibility: ∀U,V. (A U ≬ J V) =_1 (U ≬ J V)
+ }.
+
+record continuous_relation (S,T: basic_topology) : Type[1] ≝
+ { cont_rel:> S ⇒_\r1 T;
+ reduced: ∀U. U =_1 J ? U → image_coercion ?? cont_rel U =_1 J ? (image_coercion ?? cont_rel U);
+ saturated: ∀U. U =_1 A ? U → (cont_rel)⎻* U = _1A ? ((cont_rel)⎻* U)
+ }.
+
+definition continuous_relation_setoid: basic_topology → basic_topology → setoid1.
+ intros (S T); constructor 1;
+ [ apply (continuous_relation S T)
+ | constructor 1;
+ [ apply (λr,s:continuous_relation S T.∀b. A ? (ext ?? r b) = A ? (ext ?? s b));
+ | simplify; intros; apply refl1;
+ | simplify; intros (x y H); apply sym1; apply H
+ | simplify; intros; apply trans1; [2: apply f |3: apply f1; |1: skip]]]
+qed.
+
+definition continuos_relation_of_continuous_relation_setoid :
+ ∀P,Q. continuous_relation_setoid P Q → continuous_relation P Q ≝ λP,Q,x.x.
+coercion continuos_relation_of_continuous_relation_setoid.
+
+axiom continuous_relation_eq':
+ ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
+ a = a' → ∀X.a⎻* (A o1 X) = a'⎻* (A o1 X).
+(*
+ intros; split; intro; unfold minus_star_image; simplify; intros;
+ [ cut (ext ?? a a1 ⊆ A ? X); [2: intros 2; apply (H1 a2); cases f1; assumption;]
+ lapply (if ?? (A_is_saturation ???) Hcut); clear Hcut;
+ cut (A ? (ext ?? a' a1) ⊆ A ? X); [2: apply (. (H ?)‡#); assumption]
+ lapply (fi ?? (A_is_saturation ???) Hcut);
+ apply (Hletin1 x); change with (x ∈ ext ?? a' a1); split; simplify;
+ [ apply I | assumption ]
+ | cut (ext ?? a' a1 ⊆ A ? X); [2: intros 2; apply (H1 a2); cases f1; assumption;]
+ lapply (if ?? (A_is_saturation ???) Hcut); clear Hcut;
+ cut (A ? (ext ?? a a1) ⊆ A ? X); [2: apply (. (H ?)\sup -1‡#); assumption]
+ lapply (fi ?? (A_is_saturation ???) Hcut);
+ apply (Hletin1 x); change with (x ∈ ext ?? a a1); split; simplify;
+ [ apply I | assumption ]]
+qed.*)
+
+lemma continuous_relation_eq_inv':
+ ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
+ (∀X.a⎻* (A o1 X) = a'⎻* (A o1 X)) → a=a'.
+ intros 6;
+ cut (∀a,a': continuous_relation_setoid o1 o2.
+ (∀X.a⎻* (A o1 X) = a'⎻* (A o1 X)) →
+ ∀V:o2. A ? (ext ?? a' V) ⊆ A ? (ext ?? a V));
+ [2: clear b f a' a; intros;
+ lapply depth=0 (λV.saturation_expansive ??? (extS ?? a V)); [2: apply A_is_saturation;|skip]
+ (* fundamental adjunction here! to be taken out *)
+ cut (∀V:Ω^o2.V ⊆ a⎻* (A ? (extS ?? a V)));
+ [2: intro; intros 2; unfold minus_star_image; simplify; intros;
+ apply (Hletin V1 x); whd; split; [ exact I | exists; [apply a1] split; assumption]]
+ clear Hletin;
+ cut (∀V:Ω^o2.V ⊆ a'⎻* (A ? (extS ?? a V)));
+ [2: intro; apply (. #‡(f ?)^-1); apply Hcut] clear f Hcut;
+ (* second half of the fundamental adjunction here! to be taken out too *)
+ intro; lapply (Hcut1 {(V)}); clear Hcut1;
+ unfold minus_star_image in Hletin; unfold singleton in Hletin; simplify in Hletin;
+ whd in Hletin; whd in Hletin:(?→?→%); simplify in Hletin;
+ apply (if ?? (A_is_saturation ???));
+ intros 2 (x H); lapply (Hletin V ? x ?);
+ [ apply refl | unfold foo; apply H; ]
+ change with (x ∈ A ? (ext ?? a V));
+ apply (. #‡(†(extS_singleton ????)^-1));
+ assumption;]
+ split; apply Hcut; [2: assumption | intro; apply sym1; apply f]
+qed.
+
+definition continuous_relation_comp:
+ ∀o1,o2,o3.
+ continuous_relation_setoid o1 o2 →
+ continuous_relation_setoid o2 o3 →
+ continuous_relation_setoid o1 o3.
+ intros (o1 o2 o3 r s); constructor 1;
+ [ alias symbol "compose" (instance 1) = "category1 composition".
+apply (s ∘ r)
+ | intros;
+ apply sym1;
+ (*change in ⊢ (? ? ? (? ? ? ? %) ?) with (image_coercion ?? (s ∘ r) U);*)
+ apply (.= †(image_comp ??????));
+ apply (.= (reduced ?? s (image_coercion ?? r U) ?)^-1);
+ [ apply (.= (reduced ?????)); [ assumption | apply refl1 ]
+ | change in ⊢ (? ? ? % ?) with ((image_coercion ?? s ∘ image_coercion ?? r) U);
+ apply (.= (image_comp ??????)^-1);
+ apply refl1]
+ | intros;
+ apply sym1;
+ apply (.= †(minus_star_image_comp ??? s r ?));
+ apply (.= (saturated ?? s ((r)⎻* U) ?)^-1);
+ [ apply (.= (saturated ?????)); [ assumption | apply refl1 ]
+ | change in ⊢ (? ? ? % ?) with ((s⎻* ∘ r⎻* ) U);
+ apply (.= (minus_star_image_comp ??????)^-1);
+ apply refl1]]
+qed.
+
+definition BTop: category1.
+ constructor 1;
+ [ apply basic_topology
+ | apply continuous_relation_setoid
+ | intro; constructor 1;
+ [ apply id1
+ | intros;
+ apply (.= (image_id ??));
+ apply sym1;
+ apply (.= †(image_id ??));
+ apply sym1;
+ assumption
+ | intros;
+ apply (.= (minus_star_image_id ??));
+ apply sym1;
+ apply (.= †(minus_star_image_id ??));
+ apply sym1;
+ assumption]
+ | intros; constructor 1;
+ [ apply continuous_relation_comp;
+ | intros; simplify; intro x; simplify;
+ lapply depth=0 (continuous_relation_eq' ???? e) as H';
+ lapply depth=0 (continuous_relation_eq' ???? e1) as H1';
+ letin K ≝ (λX.H1' ((a)⎻* (A ? X))); clearbody K;
+ cut (∀X:Ω \sup o1.
+ (b)⎻* (A o2 ((a)⎻* (A o1 X)))
+ =_1 (b')⎻* (A o2 ((a')⎻* (A o1 X))));
+ [2: intro; apply sym1;
+ apply (.= (†(†((H' X)^-1)))); apply sym1; apply (K X);]
+ clear K H' H1';
+alias symbol "powerset" (instance 5) = "powerset low".
+alias symbol "compose" (instance 2) = "category1 composition".
+cut (∀X:Ω^o1.
+ ((b ∘ a))⎻* (A o1 X) =_1 ((b'∘a'))⎻* (A o1 X));
+ [2: intro; unfold foo;
+ apply (.= (minus_star_image_comp ??????));
+ change in ⊢ (? ? ? % ?) with ((b)⎻* ((a)⎻* (A o1 X)));
+ apply (.= †(saturated ?????));
+ [ apply ((saturation_idempotent ????)^-1); apply A_is_saturation ]
+ apply sym1;
+ apply (.= (minus_star_image_comp ??????));
+ change in ⊢ (? ? ? % ?) with ((b')⎻* ((a')⎻* (A o1 X)));
+ apply (.= †(saturated ?????));
+ [ apply ((saturation_idempotent ????)^-1); apply A_is_saturation ]
+ apply ((Hcut X)^-1)]
+ clear Hcut; generalize in match x; clear x;
+ apply (continuous_relation_eq_inv');
+ apply Hcut1;]
+ | intros; simplify; intro; do 2 (unfold continuous_relation_comp); simplify;
+ alias symbol "trans" (instance 1) = "trans1".
+alias symbol "refl" (instance 5) = "refl1".
+alias symbol "prop2" (instance 3) = "prop21".
+alias symbol "prop1" (instance 2) = "prop11".
+alias symbol "assoc" (instance 4) = "category1 assoc".
+apply (.= †(ASSOC‡#));
+ apply refl1
+ | intros; simplify; intro; unfold continuous_relation_comp; simplify;
+ apply (.= †((id_neutral_right1 ????)‡#));
+ apply refl1
+ | intros; simplify; intro; simplify;
+ apply (.= †((id_neutral_left1 ????)‡#));
+ apply refl1]
+qed.
+
+(*
+(*CSC: unused! *)
+(* this proof is more logic-oriented than set/lattice oriented *)
+theorem continuous_relation_eqS:
+ ∀o1,o2:basic_topology.∀a,a': continuous_relation_setoid o1 o2.
+ a = a' → ∀X. A ? (extS ?? a X) = A ? (extS ?? a' X).
+ intros;
+ cut (∀a: arrows1 ? o1 ?.∀x. x ∈ extS ?? a X → ∃y:o2.y ∈ X ∧ x ∈ ext ?? a y);
+ [2: intros; cases f; clear f; cases H1; exists [apply w] cases x1; split;
+ try assumption; split; assumption]
+ cut (∀a,a':continuous_relation_setoid o1 o2.eq1 ? a a' → ∀x. x ∈ extS ?? a X → ∃y:o2. y ∈ X ∧ x ∈ A ? (ext ?? a' y));
+ [2: intros; cases (Hcut ?? f); exists; [apply w] cases x1; split; try assumption;
+ apply (. #‡(H1 ?));
+ apply (saturation_expansive ?? (A_is_saturation o1) (ext ?? a1 w) x);
+ assumption;] clear Hcut;
+ split; apply (if ?? (A_is_saturation ???)); intros 2;
+ [lapply (Hcut1 a a' H a1 f) | lapply (Hcut1 a' a (H \sup -1) a1 f)]
+ cases Hletin; clear Hletin; cases x; clear x;
+ cut (∀a: arrows1 ? o1 ?. ext ?? a w ⊆ extS ?? a X);
+ [2,4: intros 3; cases f3; clear f3; simplify in f5; split; try assumption;
+ exists [1,3: apply w] split; assumption;]
+ cut (∀a. A ? (ext o1 o2 a w) ⊆ A ? (extS o1 o2 a X));
+ [2,4: intros; apply saturation_monotone; try (apply A_is_saturation); apply Hcut;]
+ apply Hcut2; assumption.
+qed.
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/basic_topologies.ma".
+include "formal_topology/o-basic_topologies.ma".
+include "formal_topology/relations_to_o-algebra.ma".
+
+definition o_basic_topology_of_basic_topology: basic_topology → Obasic_topology.
+ intros (b); constructor 1;
+ [ apply (POW' b) | apply (A b) | apply (J b);
+ | apply (A_is_saturation b) | apply (J_is_reduction b) | apply (compatibility b) ]
+qed.
+
+definition o_continuous_relation_of_continuous_relation:
+ ∀BT1,BT2.continuous_relation BT1 BT2 →
+ Ocontinuous_relation (o_basic_topology_of_basic_topology BT1) (o_basic_topology_of_basic_topology BT2).
+ intros (BT1 BT2 c); constructor 1;
+ [ apply (orelation_of_relation ?? c) | apply (reduced ?? c) | apply (saturated ?? c) ]
+qed.
+
+axiom daemon: False.
+
+lemma o_continuous_relation_of_continuous_relation_morphism :
+ ∀S,T:category2_of_category1 BTop.
+ unary_morphism2 (arrows2 (category2_of_category1 BTop) S T)
+ (arrows2 OBTop (o_basic_topology_of_basic_topology S) (o_basic_topology_of_basic_topology T)).
+intros (S T);
+ constructor 1;
+ [ apply (o_continuous_relation_of_continuous_relation S T);
+ | cases daemon (*apply (o_relation_pair_of_relation_pair_is_morphism S T)*)]
+qed.
+
+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;
+ | cases daemon (*apply o_relation_topology_of_relation_topology_morphism_respects_id*);
+ | cases daemon (*apply o_relation_topology_of_relation_topology_morphism_respects_comp*);]
+qed.
+
+theorem BTop_to_OBTop_faithful: faithful2 ?? BTop_to_OBTop.
+ intros 5; apply (continuous_relation_eq_inv' o1 o2 f g); apply e;
+qed.
+
+include "formal_topology/notation.ma".
+
+theorem BTop_to_OBTop_full: full2 ?? BTop_to_OBTop.
+ intros 3 (S T);
+ cases (POW_full (carrbt S) (carrbt T) (Ocont_rel ?? f)) (g Hg);
+ (* cases Hg; *)
+ exists [
+ constructor 1;
+ [ apply g
+ | unfold image_coercion; cases daemon (*apply hide; intros; lapply (Oreduced ?? f ? e); unfold image_coercion;
+ cases Hg; lapply (e3 U) as K; apply (.= K);
+ apply (.= Hletin); apply rule (†(K^-1)); *)
+ | cases daemon (* apply hide; intros; lapply (Osaturated ?? f ? e);
+ cases Hg; lapply (e1 U) as K; apply (.= K);
+ apply (.= Hletin); apply rule (†(K^-1)); *)
+ ]
+ | simplify; unfold BTop_to_OBTop; simplify;
+ cases Hg; unfold o_continuous_relation_of_continuous_relation_morphism;
+ simplify;
+ change with ((orelation_of_relation ?? g)⎻* ∘ oA (o_basic_topology_of_basic_topology S) =
+ f⎻* ∘ oA (o_basic_topology_of_basic_topology S));
+
+
+ change with (g⎻* ∘ oA (o_basic_topology_of_basic_topology S) =
+ f⎻* ∘ oA (o_basic_topology_of_basic_topology S));
+ apply sym2; whd in T;
+ intro;
+ apply trans2; [2: apply sym2; [2: apply Hg;
+
+ whd in ⊢ (?(??%%)???);
+ apply (.= Hg^-1);
+ unfold o_continuous_relation_of_continuous_relation_morphism; simplify;
+ intro; simplify;
+ unfold image_coercion; cases Hg; whd; simplify; intro; simplify;
+qed.
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/cprop_connectives.ma".
+
+inductive eq (A:Type[0]) (x:A) : A → CProp[0] ≝
+ refl: eq A x x.
+
+notation "hvbox(a break = \sub \ID b)" non associative with precedence 45
+for @{ 'eqID $a $b }.
+
+notation > "hvbox(a break =_\ID b)" non associative with precedence 45
+for @{ 'eqID ? $a $b }.
+
+interpretation "ID eq" 'eqID x y = (eq ? x y).
+
+record equivalence_relation (A:Type[0]) : Type[1] ≝
+ { eq_rel:2> A → A → CProp[0];
+ refl: reflexive ? eq_rel;
+ sym: symmetric ? eq_rel;
+ trans: transitive ? eq_rel
+ }.
+
+record setoid : Type[1] ≝
+ { carr:> Type[0];
+ eq: equivalence_relation carr
+ }.
+
+record equivalence_relation1 (A:Type[1]) : Type[2] ≝
+ { eq_rel1:2> A → A → CProp[1];
+ refl1: reflexive1 ? eq_rel1;
+ sym1: symmetric1 ? eq_rel1;
+ trans1: transitive1 ? eq_rel1
+ }.
+
+record setoid1: Type[2] ≝
+ { carr1:> Type[1];
+ eq1: equivalence_relation1 carr1
+ }.
+
+definition setoid1_of_setoid: setoid → setoid1.
+ #s
+ @mk_setoid1
+ [ @(carr s)
+ | @mk_equivalence_relation1
+ [ @(eq_rel s)
+ @(eq s)
+ | @(refl s)
+ | @(sym s)
+ | @(trans s)]]
+qed.
+
+coercion setoid1_of_setoid.
+(*
+prefer coercion Type_OF_setoid.
+*)
+record equivalence_relation2 (A:Type[2]) : Type[3] ≝
+ { eq_rel2:2> A → A → CProp[2];
+ refl2: reflexive2 ? eq_rel2;
+ sym2: symmetric2 ? eq_rel2;
+ trans2: transitive2 ? eq_rel2
+ }.
+
+record setoid2: Type[3] ≝
+ { carr2:> Type[2];
+ eq2: equivalence_relation2 carr2
+ }.
+
+definition setoid2_of_setoid1: setoid1 → setoid2.
+ #s
+ @mk_setoid2
+ [ @(carr1 s)
+ | @mk_equivalence_relation2
+ [ @(eq_rel1 s)
+ @(eq1 s)
+ | @(refl1 s)
+ | @(sym1 s)
+ | @(trans1 s)]]
+qed.
+
+coercion setoid2_of_setoid1.
+(*
+prefer coercion Type_OF_setoid2.
+prefer coercion Type_OF_setoid.
+prefer coercion Type_OF_setoid1.
+(* we prefer 0 < 1 < 2 *)
+*)
+record equivalence_relation3 (A:Type[3]) : Type[4] ≝
+ { eq_rel3:2> A → A → CProp[3];
+ refl3: reflexive3 ? eq_rel3;
+ sym3: symmetric3 ? eq_rel3;
+ trans3: transitive3 ? eq_rel3
+ }.
+
+record setoid3: Type[4] ≝
+ { carr3:> Type[3];
+ eq3: equivalence_relation3 carr3
+ }.
+
+interpretation "setoid3 eq" 'eq t x y = (eq_rel3 ? (eq3 t) x y).
+interpretation "setoid2 eq" 'eq t x y = (eq_rel2 ? (eq2 t) x y).
+interpretation "setoid1 eq" 'eq t x y = (eq_rel1 ? (eq1 t) x y).
+interpretation "setoid eq" 'eq t x y = (eq_rel ? (eq t) x y).
+
+notation > "hvbox(a break =_12 b)" non associative with precedence 45
+for @{ eq_rel2 (carr2 (setoid2_of_setoid1 ?)) (eq2 (setoid2_of_setoid1 ?)) $a $b }.
+notation > "hvbox(a break =_0 b)" non associative with precedence 45
+for @{ eq_rel ? (eq ?) $a $b }.
+notation > "hvbox(a break =_1 b)" non associative with precedence 45
+for @{ eq_rel1 ? (eq1 ?) $a $b }.
+notation > "hvbox(a break =_2 b)" non associative with precedence 45
+for @{ eq_rel2 ? (eq2 ?) $a $b }.
+notation > "hvbox(a break =_3 b)" non associative with precedence 45
+for @{ eq_rel3 ? (eq3 ?) $a $b }.
+
+interpretation "setoid3 symmetry" 'invert r = (sym3 ???? r).
+interpretation "setoid2 symmetry" 'invert r = (sym2 ???? r).
+interpretation "setoid1 symmetry" 'invert r = (sym1 ???? r).
+interpretation "setoid symmetry" 'invert r = (sym ???? r).
+notation ".= r" with precedence 50 for @{'trans $r}.
+interpretation "trans3" 'trans r = (trans3 ????? r).
+interpretation "trans2" 'trans r = (trans2 ????? r).
+interpretation "trans1" 'trans r = (trans1 ????? r).
+interpretation "trans" 'trans r = (trans ????? r).
+
+record unary_morphism (A,B: setoid) : Type[0] ≝
+ { fun1:(*1>*) A → B;
+ prop1: ∀a,a'. eq ? a a' → eq ? (fun1 a) (fun1 a')
+ }.
+
+record unary_morphism1 (A,B: setoid1) : Type[1] ≝
+ { fun11:1> A → B;
+ prop11: ∀a,a'. eq1 ? a a' → eq1 ? (fun11 a) (fun11 a')
+ }.
+
+record unary_morphism2 (A,B: setoid2) : Type[2] ≝
+ { fun12:1> A → B;
+ prop12: ∀a,a'. eq2 ? a a' → eq2 ? (fun12 a) (fun12 a')
+ }.
+
+record unary_morphism3 (A,B: setoid3) : Type[3] ≝
+ { fun13:1> A → B;
+ prop13: ∀a,a'. eq3 ? a a' → eq3 ? (fun13 a) (fun13 a')
+ }.
+
+record binary_morphism (A,B,C:setoid) : Type[0] ≝
+ { fun2:2> A → B → C;
+ prop2: ∀a,a',b,b'. eq ? a a' → eq ? b b' → eq ? (fun2 a b) (fun2 a' b')
+ }.
+
+record binary_morphism1 (A,B,C:setoid1) : Type[1] ≝
+ { fun21:2> A → B → C;
+ prop21: ∀a,a',b,b'. eq1 ? a a' → eq1 ? b b' → eq1 ? (fun21 a b) (fun21 a' b')
+ }.
+
+record binary_morphism2 (A,B,C:setoid2) : Type[2] ≝
+ { fun22:2> A → B → C;
+ prop22: ∀a,a',b,b'. eq2 ? a a' → eq2 ? b b' → eq2 ? (fun22 a b) (fun22 a' b')
+ }.
+
+record binary_morphism3 (A,B,C:setoid3) : Type[3] ≝
+ { fun23:2> A → B → C;
+ prop23: ∀a,a',b,b'. eq3 ? a a' → eq3 ? b b' → eq3 ? (fun23 a b) (fun23 a' b')
+ }.
+
+notation "† c" with precedence 90 for @{'prop1 $c }.
+notation "l ‡ r" with precedence 90 for @{'prop2 $l $r }.
+notation "#" with precedence 90 for @{'refl}.
+interpretation "prop1" 'prop1 c = (prop1 ????? c).
+interpretation "prop11" 'prop1 c = (prop11 ????? c).
+interpretation "prop12" 'prop1 c = (prop12 ????? c).
+interpretation "prop13" 'prop1 c = (prop13 ????? c).
+interpretation "prop2" 'prop2 l r = (prop2 ???????? l r).
+interpretation "prop21" 'prop2 l r = (prop21 ???????? l r).
+interpretation "prop22" 'prop2 l r = (prop22 ???????? l r).
+interpretation "prop23" 'prop2 l r = (prop23 ???????? l r).
+interpretation "refl" 'refl = (refl ???).
+interpretation "refl1" 'refl = (refl1 ???).
+interpretation "refl2" 'refl = (refl2 ???).
+interpretation "refl3" 'refl = (refl3 ???).
+
+notation > "A × term 74 B ⇒ term 19 C" non associative with precedence 72 for @{'binary_morphism0 $A $B $C}.
+notation > "A × term 74 B ⇒_1 term 19 C" non associative with precedence 72 for @{'binary_morphism1 $A $B $C}.
+notation > "A × term 74 B ⇒_2 term 19 C" non associative with precedence 72 for @{'binary_morphism2 $A $B $C}.
+notation > "A × term 74 B ⇒_3 term 19 C" non associative with precedence 72 for @{'binary_morphism3 $A $B $C}.
+notation > "B ⇒_1 C" right associative with precedence 72 for @{'arrows1_SET $B $C}.
+notation > "B ⇒_1. C" right associative with precedence 72 for @{'arrows1_SETlow $B $C}.
+notation > "B ⇒_2 C" right associative with precedence 72 for @{'arrows2_SET1 $B $C}.
+notation > "B ⇒_2. C" right associative with precedence 72 for @{'arrows2_SET1low $B $C}.
+
+notation "A × term 74 B ⇒ term 19 C" non associative with precedence 72 for @{'binary_morphism0 $A $B $C}.
+notation "A × term 74 B ⇒\sub 1 term 19 C" non associative with precedence 72 for @{'binary_morphism1 $A $B $C}.
+notation "A × term 74 B ⇒\sub 2 term 19 C" non associative with precedence 72 for @{'binary_morphism2 $A $B $C}.
+notation "A × term 74 B ⇒\sub 3 term 19 C" non associative with precedence 72 for @{'binary_morphism3 $A $B $C}.
+notation "B ⇒\sub 1 C" right associative with precedence 72 for @{'arrows1_SET $B $C}.
+notation "B ⇒\sub 2 C" right associative with precedence 72 for @{'arrows2_SET1 $B $C}.
+notation "B ⇒\sub 1. C" right associative with precedence 72 for @{'arrows1_SETlow $B $C}.
+notation "B ⇒\sub 2. C" right associative with precedence 72 for @{'arrows2_SET1low $B $C}.
+
+interpretation "'binary_morphism0" 'binary_morphism0 A B C = (binary_morphism A B C).
+interpretation "'arrows2_SET1 low" 'arrows2_SET1 A B = (unary_morphism2 A B).
+interpretation "'arrows2_SET1low" 'arrows2_SET1low A B = (unary_morphism2 A B).
+interpretation "'binary_morphism1" 'binary_morphism1 A B C = (binary_morphism1 A B C).
+interpretation "'binary_morphism2" 'binary_morphism2 A B C = (binary_morphism2 A B C).
+interpretation "'binary_morphism3" 'binary_morphism3 A B C = (binary_morphism3 A B C).
+interpretation "'arrows1_SET low" 'arrows1_SET A B = (unary_morphism1 A B).
+interpretation "'arrows1_SETlow" 'arrows1_SETlow A B = (unary_morphism1 A B).
+
+
+definition unary_morphism2_of_unary_morphism1:
+ ∀S,T:setoid1.unary_morphism1 S T → unary_morphism2 (setoid2_of_setoid1 S) T.
+ #S #T #u
+ @mk_unary_morphism2
+ [ @(fun11 ?? u)
+ | @(prop11 ?? u) ]
+qed.
+
+definition CPROP: setoid1.
+ @mk_setoid1
+ [ @CProp[0]
+ | @mk_equivalence_relation1
+ [ @Iff
+ | #H1 @mk_Iff #H3 #H5 assumption
+ | #H7 #H8 #i cases i #H14 #H15 @mk_Iff assumption
+ | #H17 #H18 #H19 #i1 #i cases i cases i1 #f #f1 #f2 #f3 @mk_Iff
+ [ #x1 @(f2 (f x1)) | #z1 @(f1 (f3 z1))]]]
+qed.
+
+definition CProp0_of_CPROP: carr1 CPROP → CProp[0] ≝ λx.x.
+coercion CProp0_of_CPROP.
+
+alias symbol "eq" = "setoid1 eq".
+definition fi': ∀A,B:CPROP. A = B → B → A.
+ intros; @(fi ?? e); assumption.
+qed.
+
+notation ". r" with precedence 50 for @{'fi $r}.
+interpretation "fi" 'fi r = (fi' ?? r).
+
+definition and_morphism: binary_morphism1 CPROP CPROP CPROP.
+ constructor 1;
+ [ @And
+ | intros; split; intro; cases a1; split;
+ [ @(if ?? e a2)
+ | @(if ?? e1 b1)
+ | @(fi ?? e a2)
+ | @(fi ?? e1 b1)]]
+qed.
+
+interpretation "and_morphism" 'and a b = (fun21 ??? and_morphism a b).
+
+definition or_morphism: binary_morphism1 CPROP CPROP CPROP.
+ constructor 1;
+ [ @Or
+ | intros; split; intro; cases o; [1,3:left |2,4: right]
+ [ @(if ?? e a1)
+ | @(fi ?? e a1)
+ | @(if ?? e1 b1)
+ | @(fi ?? e1 b1)]]
+qed.
+
+interpretation "or_morphism" 'or a b = (fun21 ??? or_morphism a b).
+
+definition if_morphism: binary_morphism1 CPROP CPROP CPROP.
+ constructor 1;
+ [ @(λA,B. A → B)
+ | intros; split; intros;
+ [ @(if ?? e1); @f; @(fi ?? e); assumption
+ | @(fi ?? e1); @f; @(if ?? e); assumption]]
+qed.
+
+notation > "hvbox(a break ∘ b)" left associative with precedence 55 for @{ comp ??? $a $b }.
+record category : Type[1] ≝ {
+ objs:> Type[0];
+ arrows: objs → objs → setoid;
+ id: ∀o:objs. arrows o o;
+ comp: ∀o1,o2,o3. (arrows o1 o2) × (arrows o2 o3) ⇒ (arrows o1 o3);
+ comp_assoc: ∀o1,o2,o3,o4.∀a12:arrows o1 ?.∀a23:arrows o2 ?.∀a34:arrows o3 o4.
+ (a12 ∘ a23) ∘ a34 =_0 a12 ∘ (a23 ∘ a34);
+ id_neutral_left : ∀o1,o2. ∀a: arrows o1 o2. (id o1) ∘ a =_0 a;
+ id_neutral_right: ∀o1,o2. ∀a: arrows o1 o2. a ∘ (id o2) =_0 a
+}.
+notation "hvbox(a break ∘ b)" left associative with precedence 55 for @{ 'compose $a $b }.
+
+record category1 : Type[2] ≝
+ { objs1:> Type[1];
+ arrows1: objs1 → objs1 → setoid1;
+ id1: ∀o:objs1. arrows1 o o;
+ comp1: ∀o1,o2,o3. binary_morphism1 (arrows1 o1 o2) (arrows1 o2 o3) (arrows1 o1 o3);
+ comp_assoc1: ∀o1,o2,o3,o4. ∀a12,a23,a34.
+ comp1 o1 o3 o4 (comp1 o1 o2 o3 a12 a23) a34 =_1 comp1 o1 o2 o4 a12 (comp1 o2 o3 o4 a23 a34);
+ id_neutral_right1: ∀o1,o2. ∀a: arrows1 o1 o2. comp1 ??? (id1 o1) a =_1 a;
+ id_neutral_left1: ∀o1,o2. ∀a: arrows1 o1 o2. comp1 ??? a (id1 o2) =_1 a
+ }.
+
+record category2 : Type[3] ≝
+ { objs2:> Type[2];
+ arrows2: objs2 → objs2 → setoid2;
+ id2: ∀o:objs2. arrows2 o o;
+ comp2: ∀o1,o2,o3. binary_morphism2 (arrows2 o1 o2) (arrows2 o2 o3) (arrows2 o1 o3);
+ comp_assoc2: ∀o1,o2,o3,o4. ∀a12,a23,a34.
+ comp2 o1 o3 o4 (comp2 o1 o2 o3 a12 a23) a34 =_2 comp2 o1 o2 o4 a12 (comp2 o2 o3 o4 a23 a34);
+ id_neutral_right2: ∀o1,o2. ∀a: arrows2 o1 o2. comp2 ??? (id2 o1) a =_2 a;
+ id_neutral_left2: ∀o1,o2. ∀a: arrows2 o1 o2. comp2 ??? a (id2 o2) =_2 a
+ }.
+
+record category3 : Type[4] ≝
+ { objs3:> Type[3];
+ arrows3: objs3 → objs3 → setoid3;
+ id3: ∀o:objs3. arrows3 o o;
+ comp3: ∀o1,o2,o3. binary_morphism3 (arrows3 o1 o2) (arrows3 o2 o3) (arrows3 o1 o3);
+ comp_assoc3: ∀o1,o2,o3,o4. ∀a12,a23,a34.
+ comp3 o1 o3 o4 (comp3 o1 o2 o3 a12 a23) a34 =_3 comp3 o1 o2 o4 a12 (comp3 o2 o3 o4 a23 a34);
+ id_neutral_right3: ∀o1,o2. ∀a: arrows3 o1 o2. comp3 ??? (id3 o1) a =_3 a;
+ id_neutral_left3: ∀o1,o2. ∀a: arrows3 o1 o2. comp3 ??? a (id3 o2) =_3 a
+ }.
+
+notation "'ASSOC'" with precedence 90 for @{'assoc}.
+
+interpretation "category2 composition" 'compose x y = (fun22 ??? (comp2 ????) y x).
+interpretation "category2 assoc" 'assoc = (comp_assoc2 ????????).
+interpretation "category1 composition" 'compose x y = (fun21 ??? (comp1 ????) y x).
+interpretation "category1 assoc" 'assoc = (comp_assoc1 ????????).
+interpretation "category composition" 'compose x y = (fun2 ??? (comp ????) y x).
+interpretation "category assoc" 'assoc = (comp_assoc ????????).
+
+definition category2_of_category1: category1 → category2.
+ intro;
+ constructor 1;
+ [ @(objs1 c);
+ | intros; @(setoid2_of_setoid1 (arrows1 c o o1));
+ | @(id1 c);
+ | intros;
+ constructor 1;
+ [ intros; @(comp1 c o1 o2 o3 c1 c2);
+ | intros; unfold setoid2_of_setoid1 in e e1 a a' b b'; simplify in e e1 a a' b b';
+ change with ((b∘a) =_1 (b'∘a')); @(e‡e1); ]
+ | intros; simplify; whd in a12 a23 a34; whd; @rule (ASSOC);
+ | intros; simplify; whd in a; whd; @id_neutral_right1;
+ | intros; simplify; whd in a; whd; @id_neutral_left1; ]
+qed.
+(*coercion category2_of_category1.*)
+
+record functor2 (C1: category2) (C2: category2) : Type[3] ≝
+ { map_objs2:1> C1 → C2;
+ map_arrows2: ∀S,T. unary_morphism2 (arrows2 ? S T) (arrows2 ? (map_objs2 S) (map_objs2 T));
+ respects_id2: ∀o:C1. map_arrows2 ?? (id2 ? o) = id2 ? (map_objs2 o);
+ respects_comp2:
+ ∀o1,o2,o3.∀f1:arrows2 ? o1 o2.∀f2:arrows2 ? o2 o3.
+ map_arrows2 ?? (f2 ∘ f1) = map_arrows2 ?? f2 ∘ map_arrows2 ?? f1}.
+
+notation > "F⎽⇒ x" left associative with precedence 60 for @{'map_arrows2 $F $x}.
+notation "F\sub⇒ x" left associative with precedence 60 for @{'map_arrows2 $F $x}.
+interpretation "map_arrows2" 'map_arrows2 F x = (fun12 ?? (map_arrows2 ?? F ??) x).
+
+definition functor2_setoid: category2 → category2 → setoid3.
+ intros (C1 C2);
+ constructor 1;
+ [ @(functor2 C1 C2);
+ | constructor 1;
+ [ intros (f g);
+ @(∀c:C1. cic:/matita/logic/equality/eq.ind#xpointer(1/1) ? (f c) (g c));
+ | simplify; intros; @cic:/matita/logic/equality/eq.ind#xpointer(1/1/1);
+ | simplify; intros; @cic:/matita/logic/equality/sym_eq.con; @H;
+ | simplify; intros; @cic:/matita/logic/equality/trans_eq.con;
+ [2: @H; | skip | @H1;]]]
+qed.
+
+definition functor2_of_functor2_setoid: ∀S,T. functor2_setoid S T → functor2 S T ≝ λS,T,x.x.
+coercion functor2_of_functor2_setoid.
+
+definition CAT2: category3.
+ constructor 1;
+ [ @category2;
+ | @functor2_setoid;
+ | intros; constructor 1;
+ [ @(λx.x);
+ | intros; constructor 1;
+ [ @(λx.x);
+ | intros; assumption;]
+ | intros; @rule #;
+ | intros; @rule #; ]
+ | intros; constructor 1;
+ [ intros; constructor 1;
+ [ intros; @(c1 (c o));
+ | intros; constructor 1;
+ [ intro; @(map_arrows2 ?? c1 ?? (map_arrows2 ?? c ?? c2));
+ | intros; @(††e); ]
+ | intros; simplify;
+ @(.= †(respects_id2 : ?));
+ @(respects_id2 : ?);
+ | intros; simplify;
+ @(.= †(respects_comp2 : ?));
+ @(respects_comp2 : ?); ]
+ | intros; intro; simplify;
+ @(cic:/matita/logic/equality/eq_ind.con ????? (e ?));
+ @(cic:/matita/logic/equality/eq_ind.con ????? (e1 ?));
+ constructor 1; ]
+ | intros; intro; simplify; constructor 1;
+ | intros; intro; simplify; constructor 1;
+ | intros; intro; simplify; constructor 1; ]
+qed.
+
+definition category2_of_objs3_CAT2: objs3 CAT2 → category2 ≝ λx.x.
+coercion category2_of_objs3_CAT2.
+
+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;
+ [ @(unary_morphism s s1);
+ | constructor 1;
+ [ intros (f g); @(∀a:s. eq ? (f a) (g a));
+ | intros 1; simplify; intros; @refl;
+ | simplify; intros; @sym; @f;
+ | simplify; intros; @trans; [2: @f; | skip | @f1]]]
+qed.
+
+definition SET: category1.
+ constructor 1;
+ [ @setoid;
+ | @rule (λS,T:setoid.setoid1_of_setoid (unary_morphism_setoid S T));
+ | intros; constructor 1; [ @(λx:carr o.x); | intros; assumption ]
+ | intros; constructor 1; [ intros; constructor 1; [ @(λx. c1 (c x)); | intros;
+ @(† (†e));]
+ | intros; whd; intros; simplify; whd in H1; whd in H;
+ @trans; [ @(b (a' a1)); | l@(prop1 ?? b (a a1) (a' a1));
+ [ @Hletin | @(e a1); ] | @e1; ]]
+ | intros; whd; intros; simplify; @refl;
+ | intros; simplify; whd; intros; simplify; @refl;
+ | intros; simplify; whd; intros; simplify; @refl;
+ ]
+qed.
+
+definition setoid_of_SET: objs1 SET → setoid ≝ λx.x.
+coercion setoid_of_SET.
+
+definition unary_morphism_setoid_of_arrows1_SET:
+ ∀P,Q.arrows1 SET P Q → unary_morphism_setoid P Q ≝ λP,Q,x.x.
+coercion unary_morphism_setoid_of_arrows1_SET.
+
+interpretation "'arrows1_SET" 'arrows1_SET A B = (arrows1 SET A B).
+
+definition unary_morphism1_setoid1: setoid1 → setoid1 → setoid1.
+ intros;
+ constructor 1;
+ [ @(unary_morphism1 s s1);
+ | constructor 1;
+ [ intros (f g);
+ alias symbol "eq" = "setoid1 eq".
+ @(∀a: carr1 s. f a = g a);
+ | intros 1; simplify; intros; @refl1;
+ | simplify; intros; @sym1; @f;
+ | simplify; intros; @trans1; [2: @f; | skip | @f1]]]
+qed.
+
+definition unary_morphism1_of_unary_morphism1_setoid1 :
+ ∀S,T. unary_morphism1_setoid1 S T → unary_morphism1 S T ≝ λP,Q,x.x.
+coercion unary_morphism1_of_unary_morphism1_setoid1.
+
+definition SET1: objs3 CAT2.
+ constructor 1;
+ [ @setoid1;
+ | @rule (λS,T.setoid2_of_setoid1 (unary_morphism1_setoid1 S T));
+ | intros; constructor 1; [ @(λx.x); | intros; assumption ]
+ | intros; constructor 1; [ intros; constructor 1; [ @(λx. c1 (c x)); | intros;
+ @(† (†e));]
+ | intros; whd; intros; simplify; whd in H1; whd in H;
+ @trans1; [ @(b (a' a1)); | l@(prop11 ?? b (a a1) (a' a1));
+ [ @Hletin | @(e a1); ] | @e1; ]]
+ | intros; whd; intros; simplify; @refl1;
+ | intros; simplify; whd; intros; simplify; @refl1;
+ | intros; simplify; whd; intros; simplify; @refl1;
+ ]
+qed.
+
+interpretation "'arrows2_SET1" 'arrows2_SET1 A B = (arrows2 SET1 A B).
+
+definition setoid1_of_SET1: objs2 SET1 → setoid1 ≝ λx.x.
+coercion setoid1_of_SET1.
+
+definition unary_morphism1_setoid1_of_arrows2_SET1:
+ ∀P,Q.arrows2 SET1 P Q → unary_morphism1_setoid1 P Q ≝ λP,Q,x.x.
+coercion unary_morphism1_setoid1_of_arrows2_SET1.
+
+variant objs2_of_category1: objs1 SET → objs2 SET1 ≝ setoid1_of_setoid.
+coercion objs2_of_category1.
+
+prefer coercion Type_OF_setoid. (* we prefer the lower carrier projection *)
+prefer coercion Type_OF_objs1.
+
+alias symbol "exists" (instance 1) = "CProp[2] 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.
+
+
+notation "r \sup *" non associative with precedence 90 for @{'OR_f_star $r}.
+notation > "r *" non associative with precedence 90 for @{'OR_f_star $r}.
+
+notation "r \sup (⎻* )" non associative with precedence 90 for @{'OR_f_minus_star $r}.
+notation > "r⎻*" non associative with precedence 90 for @{'OR_f_minus_star $r}.
+
+notation "r \sup ⎻" non associative with precedence 90 for @{'OR_f_minus $r}.
+notation > "r⎻" non associative with precedence 90 for @{'OR_f_minus $r}.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/basic_pairs.ma".
+
+(* carr1 e' necessario perche' ci sega via la coercion per gli oggetti di REL!
+ (confondendola con la coercion per gli oggetti di SET
+record concrete_space : Type[1] ≝
+ { bp:> BP;
+ converges: ∀a: carr1 (concr bp).∀U,V: carr1 (form bp). a ⊩ U → a ⊩ V → a ⊩ (U ↓ V);
+ all_covered: ∀x: carr1 (concr bp). x ⊩ form bp
+ }.
+
+record convergent_relation_pair (CS1,CS2: concrete_space) : Type[1] ≝
+ { rp:> arrows1 ? CS1 CS2;
+ respects_converges:
+ ∀b,c.
+ minus_image ?? rp \sub\c (BPextS CS2 (b ↓ c)) =
+ BPextS CS1 ((minus_image ?? rp \sub\f b) ↓ (minus_image ?? rp \sub\f c));
+ respects_all_covered:
+ minus_image ?? rp\sub\c (BPextS CS2 (full_subset (form CS2))) = BPextS CS1 (full_subset (form CS1))
+ }.
+
+definition convergent_relation_space_setoid: concrete_space → concrete_space → setoid1.
+ intros;
+ constructor 1;
+ [ apply (convergent_relation_pair c c1)
+ | constructor 1;
+ [ intros;
+ apply (relation_pair_equality c c1 c2 c3);
+ | intros 1; apply refl1;
+ | intros 2; apply sym1;
+ | intros 3; apply trans1]]
+qed.
+
+definition convergent_relation_space_composition:
+ ∀o1,o2,o3: concrete_space.
+ binary_morphism1
+ (convergent_relation_space_setoid o1 o2)
+ (convergent_relation_space_setoid o2 o3)
+ (convergent_relation_space_setoid o1 o3).
+ intros; constructor 1;
+ [ intros; whd in c c1 ⊢ %;
+ constructor 1;
+ [ apply (fun1 ??? (comp1 BP ???)); [apply (bp o2) |*: apply rp; assumption]
+ | intros;
+ change in ⊢ (? ? ? (? ? ? (? ? ? %) ?) ?) with (c1 \sub \c ∘ c \sub \c);
+ change in ⊢ (? ? ? ? (? ? ? ? (? ? ? ? ? (? ? ? (? ? ? %) ?) ?)))
+ with (c1 \sub \f ∘ c \sub \f);
+ change in ⊢ (? ? ? ? (? ? ? ? (? ? ? ? ? ? (? ? ? (? ? ? %) ?))))
+ with (c1 \sub \f ∘ c \sub \f);
+ apply (.= (extS_com ??????));
+ apply (.= (†(respects_converges ?????)));
+ apply (.= (respects_converges ?????));
+ apply (.= (†(((extS_com ??????) \sup -1)‡(extS_com ??????)\sup -1)));
+ apply refl1;
+ | change in ⊢ (? ? ? (? ? ? (? ? ? %) ?) ?) with (c1 \sub \c ∘ c \sub \c);
+ apply (.= (extS_com ??????));
+ apply (.= (†(respects_all_covered ???)));
+ apply (.= respects_all_covered ???);
+ apply refl1]
+ | intros;
+ change with (b ∘ a = b' ∘ a');
+ change in H with (rp'' ?? a = rp'' ?? a');
+ change in H1 with (rp'' ?? b = rp ?? b');
+ apply (.= (H‡H1));
+ apply refl1]
+qed.
+
+definition CSPA: category1.
+ constructor 1;
+ [ apply concrete_space
+ | apply convergent_relation_space_setoid
+ | intro; constructor 1;
+ [ apply id1
+ | intros;
+ unfold id; simplify;
+ apply (.= (equalset_extS_id_X_X ??));
+ apply (.= (†((equalset_extS_id_X_X ??)\sup -1‡
+ (equalset_extS_id_X_X ??)\sup -1)));
+ apply refl1;
+ | apply (.= (equalset_extS_id_X_X ??));
+ apply refl1]
+ | apply convergent_relation_space_composition
+ | intros; simplify;
+ change with (a34 ∘ (a23 ∘ a12) = (a34 ∘ a23) ∘ a12);
+ apply (.= ASSOC1);
+ apply refl1
+ | intros; simplify;
+ change with (a ∘ id1 ? o1 = a);
+ apply (.= id_neutral_right1 ????);
+ apply refl1
+ | intros; simplify;
+ change with (id1 ? o2 ∘ a = a);
+ apply (.= id_neutral_left1 ????);
+ apply refl1]
+qed.
+*)
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/concrete_spaces.ma".
+include "formal_topology/o-concrete_spaces.ma".
+include "formal_topology/basic_pairs_to_o-basic_pairs.ma".
+
+(*
+(* Qui, per fare le cose per bene, ci serve la nozione di funtore categorico *)
+definition o_concrete_space_of_concrete_space: cic:/matita/formal_topology/concrete_spaces/concrete_space.ind#xpointer(1/1) → concrete_space.
+ intro;
+ constructor 1;
+ [ apply (o_basic_pair_of_basic_pair (bp c));
+ | lapply depth=0 (downarrow c);
+ constructor 1;
+ [ apply
+ |
+ apply (o_operator_of_operator);fintersectsS);
+ |
+ |
+ |
+ |
+ |
+ ]
+qed.
+
+definition o_convergent_relation_pair_of_convergent_relation_pair:
+ ∀BP1,BP2.cic:/matita/formal_topology/concrete_spaces/convergent_relation_pair.ind#xpointer(1/1) BP1 BP2 →
+ convergent_relation_pair (o_concrete_space_of_concrete_space BP1) (o_concrete_space_of_concrete_space BP2).
+ intros;
+ constructor 1;
+ [ apply (orelation_of_relation ?? (r \sub \c));
+ | apply (orelation_of_relation ?? (r \sub \f));
+ | lapply (commute ?? r);
+ lapply (orelation_of_relation_preserves_equality ???? Hletin);
+ apply (.= (orelation_of_relation_preserves_composition (concr BP1) ??? (rel BP2)) ^ -1);
+ apply (.= (orelation_of_relation_preserves_equality ???? (commute ?? r)));
+ apply (orelation_of_relation_preserves_composition ?? (form BP2) (rel BP1) ?); ]
+qed.
+
+*)
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "basics/pts.ma".
+
+inductive False: CProp[0] ≝.
+
+inductive Or (A,B:CProp[0]) : CProp[0] ≝
+ | Left : A → Or A B
+ | Right : B → Or A B.
+
+interpretation "constructive or" 'or x y = (Or x y).
+
+inductive Or3 (A,B,C:CProp[0]) : CProp[0] ≝
+ | Left3 : A → Or3 A B C
+ | Middle3 : B → Or3 A B C
+ | Right3 : C → Or3 A B C.
+
+interpretation "constructive ternary or" 'or3 x y z= (Or3 x y z).
+
+notation < "hvbox(a break ∨ b break ∨ c)" with precedence 35 for @{'or3 $a $b $c}.
+
+inductive Or4 (A,B,C,D:CProp[0]) : CProp[0] ≝
+ | Left3 : A → Or4 A B C D
+ | Middle3 : B → Or4 A B C D
+ | Right3 : C → Or4 A B C D
+ | Extra3: D → Or4 A B C D.
+
+interpretation "constructive ternary or" 'or4 x y z t = (Or4 x y z t).
+
+notation < "hvbox(a break ∨ b break ∨ c break ∨ d)" with precedence 35 for @{'or4 $a $b $c $d}.
+
+inductive And (A,B:CProp[0]) : CProp[0] ≝
+ | Conj : A → B → And A B.
+
+interpretation "constructive and" 'and x y = (And x y).
+
+inductive And3 (A,B,C:CProp[0]) : CProp[0] ≝
+ | Conj3 : A → B → C → And3 A B C.
+
+notation < "hvbox(a break ∧ b break ∧ c)" with precedence 35 for @{'and3 $a $b $c}.
+
+interpretation "constructive ternary and" 'and3 x y z = (And3 x y z).
+
+inductive And42 (A,B,C,D:CProp[2]) : CProp[2] ≝
+ | Conj42 : A → B → C → D → And42 A B C D.
+
+notation < "hvbox(a break ∧ b break ∧ c break ∧ d)" with precedence 35 for @{'and4 $a $b $c $d}.
+
+interpretation "constructive quaternary and2" 'and4 x y z t = (And42 x y z t).
+
+record Iff (A,B:CProp[0]) : CProp[0] ≝
+ { if: A → B;
+ fi: B → A
+ }.
+
+record Iff1 (A,B:CProp[1]) : CProp[1] ≝
+ { if1: A → B;
+ fi1: B → A
+ }.
+
+notation "hvbox(a break ⇔ b)" right associative with precedence 25 for @{'iff1 $a $b}.
+interpretation "logical iff" 'iff x y = (Iff x y).
+interpretation "logical iff type1" 'iff1 x y = (Iff1 x y).
+
+inductive exT22 (A:Type[2]) (P:A→CProp[2]) : CProp[2] ≝
+ ex_introT22: ∀w:A. P w → exT22 A P.
+
+interpretation "CProp[2] exists" 'exists \eta.x = (exT22 ? x).
+(*
+definition pi1exT22 ≝ λA,P.λx:exT22 A P.match x with [ex_introT22 x _ ⇒ x].
+definition pi2exT22 ≝
+ λA,P.λx:exT22 A P.match x return λx.P (pi1exT22 ?? x) with [ex_introT22 _ p ⇒ p].
+
+interpretation "exT22 \fst" 'pi1 = (pi1exT22 ? ?).
+interpretation "exT22 \snd" 'pi2 = (pi2exT22 ? ?).
+interpretation "exT22 \fst a" 'pi1a x = (pi1exT22 ? ? x).
+interpretation "exT22 \snd a" 'pi2a x = (pi2exT22 ? ? x).
+interpretation "exT22 \fst b" 'pi1b x y = (pi1exT22 ? ? x y).
+interpretation "exT22 \snd b" 'pi2b x y = (pi2exT22 ? ? x y).
+*)
+inductive exT (A:Type[0]) (P:A→CProp[0]) : CProp[0] ≝
+ ex_introT: ∀w:A. P w → exT A P.
+
+interpretation "CProp exists" 'exists \eta.x = (exT ? x).
+
+notation "\ll term 19 a, break term 19 b \gg"
+with precedence 90 for @{'dependent_pair $a $b}.
+interpretation "dependent pair" 'dependent_pair a b =
+ (ex_introT ? ? a b).
+
+(*
+definition pi1exT ≝ λA,P.λx:exT A P.match x with [ex_introT x _ ⇒ x].
+definition pi2exT ≝
+ λA,P.λx:exT A P.match x return λx.P (pi1exT ?? x) with [ex_introT _ p ⇒ p].
+
+interpretation "exT \fst" 'pi1 = (pi1exT ? ?).
+interpretation "exT \fst a" 'pi1a x = (pi1exT ? ? x).
+interpretation "exT \fst b" 'pi1b x y = (pi1exT ? ? x y).
+interpretation "exT \snd" 'pi2 = (pi2exT ? ?).
+interpretation "exT \snd a" 'pi2a x = (pi2exT ? ? x).
+interpretation "exT \snd b" 'pi2b x y = (pi2exT ? ? x y).
+*)
+inductive exT23 (A:Type[0]) (P:A→CProp[0]) (Q:A→CProp[0]) (R:A→A→CProp[0]) : CProp[0] ≝
+ ex_introT23: ∀w,p:A. P w → Q p → R w p → exT23 A P Q R.
+(*
+definition pi1exT23 ≝
+ λA,P,Q,R.λx:exT23 A P Q R.match x with [ex_introT23 x _ _ _ _ ⇒ x].
+definition pi2exT23 ≝
+ λA,P,Q,R.λx:exT23 A P Q R.match x with [ex_introT23 _ x _ _ _ ⇒ x].
+
+interpretation "exT2 \fst" 'pi1 = (pi1exT23 ? ? ? ?).
+interpretation "exT2 \snd" 'pi2 = (pi2exT23 ? ? ? ?).
+interpretation "exT2 \fst a" 'pi1a x = (pi1exT23 ? ? ? ? x).
+interpretation "exT2 \snd a" 'pi2a x = (pi2exT23 ? ? ? ? x).
+interpretation "exT2 \fst b" 'pi1b x y = (pi1exT23 ? ? ? ? x y).
+interpretation "exT2 \snd b" 'pi2b x y = (pi2exT23 ? ? ? ? x y).
+*)
+inductive exT2 (A:Type[0]) (P,Q:A→CProp[0]) : CProp[0] ≝
+ ex_introT2: ∀w:A. P w → Q w → exT2 A P Q.
+
+
+definition Not : CProp[0] → CProp[0] ≝ λx:CProp[0].x → False.
+
+interpretation "constructive not" 'not x = (Not x).
+
+definition cotransitive: ∀C:Type[0]. ∀lt:C→C→CProp[0].CProp[0] ≝
+ λC:Type[0].λlt:C→C→CProp[0].∀x,y,z:C. lt x y → lt x z ∨ lt z y.
+
+definition coreflexive: ∀C:Type[0]. ∀lt:C→C→CProp[0].CProp[0] ≝
+ λC:Type[0].λlt:C→C→CProp[0]. ∀x:C. ¬ (lt x x).
+
+definition symmetric: ∀C:Type[0]. ∀lt:C→C→CProp[0].CProp[0] ≝
+ λC:Type[0].λlt:C→C→CProp[0]. ∀x,y:C.lt x y → lt y x.
+
+definition antisymmetric: ∀A:Type[0]. ∀R:A→A→CProp[0]. ∀eq:A→A→Prop.CProp[0] ≝
+ λA:Type[0].λR:A→A→CProp[0].λeq:A→A→Prop.∀x:A.∀y:A.R x y→R y x→eq x y.
+
+definition reflexive: ∀C:Type[0]. ∀lt:C→C→CProp[0].CProp[0] ≝ λA:Type[0].λR:A→A→CProp[0].∀x:A.R x x.
+
+definition transitive: ∀C:Type[0]. ∀lt:C→C→CProp[0].CProp[0] ≝ λA:Type[0].λR:A→A→CProp[0].∀x,y,z:A.R x y → R y z → R x z.
+
+definition reflexive1: ∀A:Type[1].∀R:A→A→CProp[1].CProp[1] ≝ λA:Type[1].λR:A→A→CProp[1].∀x:A.R x x.
+definition symmetric1: ∀A:Type[1].∀R:A→A→CProp[1].CProp[1] ≝ λC:Type[1].λlt:C→C→CProp[1]. ∀x,y:C.lt x y → lt y x.
+definition transitive1: ∀A:Type[1].∀R:A→A→CProp[1].CProp[1] ≝ λA:Type[1].λR:A→A→CProp[1].∀x,y,z:A.R x y → R y z → R x z.
+
+definition reflexive2: ∀A:Type[2].∀R:A→A→CProp[2].CProp[2] ≝ λA:Type[2].λR:A→A→CProp[2].∀x:A.R x x.
+definition symmetric2: ∀A:Type[2].∀R:A→A→CProp[2].CProp[2] ≝ λC:Type[2].λlt:C→C→CProp[2]. ∀x,y:C.lt x y → lt y x.
+definition transitive2: ∀A:Type[2].∀R:A→A→CProp[2].CProp[2] ≝ λA:Type[2].λR:A→A→CProp[2].∀x,y,z:A.R x y → R y z → R x z.
+
+definition reflexive3: ∀A:Type[3].∀R:A→A→CProp[3].CProp[3] ≝ λA:Type[3].λR:A→A→CProp[3].∀x:A.R x x.
+definition symmetric3: ∀A:Type[3].∀R:A→A→CProp[3].CProp[3] ≝ λC:Type[3].λlt:C→C→CProp[3]. ∀x,y:C.lt x y → lt y x.
+definition transitive3: ∀A:Type[3].∀R:A→A→CProp[3].CProp[3] ≝ λA:Type[3].λR:A→A→CProp[3].∀x,y,z:A.R x y → R y z → R x z.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/basic_topologies.ma".
+
+(*
+definition downarrow: ∀S:BTop. unary_morphism (Ω \sup S) (Ω \sup S).
+ intros; constructor 1;
+ [ apply (λU:Ω \sup S.{a | ∃b:carrbt S. b ∈ U ∧ a ∈ A ? (singleton ? b)});
+ intros; simplify; split; intro; cases H1; cases x; exists [1,3: apply w]
+ split; try assumption; [ apply (. H‡#) | apply (. H \sup -1‡#) ] assumption
+ | intros; split; intros 2; cases f; exists; [1,3: apply w] cases x; split;
+ try assumption; [ apply (. #‡H) | apply (. #‡H \sup -1)] assumption]
+qed.
+
+interpretation "downarrow" 'downarrow a = (fun_1 ?? (downarrow ?) a).
+
+definition ffintersects: ∀S:BTop. binary_morphism1 (Ω \sup S) (Ω \sup S) (Ω \sup S).
+ intros; constructor 1;
+ [ apply (λU,V. ↓U ∩ ↓V);
+ | intros; apply (.= (†H)‡(†H1)); apply refl1]
+qed.
+
+interpretation "ffintersects" 'fintersects U V = (fun1 ??? (ffintersects ?) U V).
+
+record formal_topology: Type ≝
+ { bt:> BTop;
+ converges: ∀U,V: Ω \sup bt. A ? (U ↓ V) = A ? U ∩ A ? V
+ }.
+
+
+definition ffintersects': ∀S:BTop. binary_morphism1 S S (Ω \sup S).
+ intros; constructor 1;
+ [ apply (λb,c:S. (singleton S b) ↓ (singleton S c));
+ | intros; apply (.= (†H)‡(†H1)); apply refl1]
+qed.
+
+interpretation "ffintersects'" 'fintersects U V = (fun1 ??? (ffintersects' ?) U V).
+
+record formal_map (S,T: formal_topology) : Type ≝
+ { cr:> continuous_relation_setoid S T;
+ C1: ∀b,c. extS ?? cr (b ↓ c) = ext ?? cr b ↓ ext ?? cr c;
+ C2: extS ?? cr T = S
+ }.
+
+definition formal_map_setoid: formal_topology → formal_topology → setoid1.
+ intros (S T); constructor 1;
+ [ apply (formal_map S T);
+ | constructor 1;
+ [ apply (λf,f1: formal_map S T.f=f1);
+ | simplify; intros 1; apply refl1
+ | simplify; intros 2; apply sym1
+ | simplify; intros 3; apply trans1]]
+qed.
+
+axiom C1':
+ ∀S,T: formal_topology.∀f:formal_map_setoid S T.∀U,V: Ω \sup T.
+ extS ?? f (U ↓ V) = extS ?? f U ↓ extS ?? f V.
+
+definition formal_map_composition:
+ ∀o1,o2,o3: formal_topology.
+ binary_morphism1
+ (formal_map_setoid o1 o2)
+ (formal_map_setoid o2 o3)
+ (formal_map_setoid o1 o3).
+ intros; constructor 1;
+ [ intros; whd in c c1; constructor 1;
+ [ apply (comp1 BTop ??? c c1);
+ | intros;
+ apply (.= (extS_com ??? c c1 ?));
+ apply (.= †(C1 ?????));
+ apply (.= (C1' ?????));
+ apply (.= ((†((extS_singleton ????) \sup -1))‡(†((extS_singleton ????) \sup -1))));
+ apply (.= (extS_com ??????)\sup -1 ‡ (extS_com ??????) \sup -1);
+ apply (.= (extS_singleton ????)‡(extS_singleton ????));
+ apply refl1;
+ | apply (.= (extS_com ??? c c1 ?));
+ apply (.= (†(C2 ???)));
+ apply (.= (C2 ???));
+ apply refl1;]
+ | intros; simplify;
+ change with (comp1 BTop ??? a b = comp1 BTop ??? a' b');
+ apply prop1; assumption]
+qed.
+
+*)
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+notation "hvbox (r \sub \c)" with precedence 90 for @{'concr_rel $r}.
+notation "hvbox (r \sub \f)" with precedence 90 for @{'form_rel $r}.
+
+definition hide ≝ λA:Type.λx:A.x.
+
+interpretation "hide" 'hide x = (hide ? x).
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/categories.ma".
+
+inductive bool : Type[0] := true : bool | false : bool.
+
+lemma BOOL : objs1 SET.
+constructor 1; [apply bool] constructor 1;
+[ intros (x y); apply (match x with [ true ⇒ match y with [ true ⇒ True | _ ⇒ False] | false ⇒ match y with [ true ⇒ False | false ⇒ True ]]);
+| whd; simplify; intros; cases x; apply I;
+| whd; simplify; intros 2; cases x; cases y; simplify; intros; assumption;
+| whd; simplify; intros 3; cases x; cases y; cases z; simplify; intros;
+ try assumption; apply I]
+qed.
+
+lemma IF_THEN_ELSE_p :
+ ∀S:setoid1.∀a,b:S.∀x,y:BOOL.x = y →
+ (λm.match m with [ true ⇒ a | false ⇒ b ]) x =
+ (λm.match m with [ true ⇒ a | false ⇒ b ]) y.
+whd in ⊢ (?→?→?→%→?);
+intros; cases x in e; cases y; simplify; intros; try apply refl1; whd in e; cases e;
+qed.
+
+interpretation "unary morphism comprehension with no proof" 'comprehension T P =
+ (mk_unary_morphism T ? P ?).
+interpretation "unary morphism1 comprehension with no proof" 'comprehension T P =
+ (mk_unary_morphism1 T ? P ?).
+
+notation > "hvbox({ ident i ∈ s | term 19 p | by })" with precedence 90
+for @{ 'comprehension_by $s (λ${ident i}. $p) $by}.
+notation < "hvbox({ ident i ∈ s | term 19 p })" with precedence 90
+for @{ 'comprehension_by $s (λ${ident i}:$_. $p) $by}.
+
+interpretation "unary morphism comprehension with proof" 'comprehension_by s \eta.f p =
+ (mk_unary_morphism s ? f p).
+interpretation "unary morphism1 comprehension with proof" 'comprehension_by s \eta.f p =
+ (mk_unary_morphism1 s ? f p).
+
+(* per il set-indexing vedere capitolo BPTools (foundational tools), Sect. 0.3.4 complete
+ lattices, Definizione 0.9 *)
+(* USARE L'ESISTENZIALE DEBOLE *)
+
+definition if_then_else ≝ λT:Type.λe,t,f.match e return λ_.T with [ true ⇒ t | false ⇒ f].
+notation > "'If' term 19 e 'then' term 19 t 'else' term 90 f" non associative with precedence 19 for @{ 'if_then_else $e $t $f }.
+notation < "'If' \nbsp term 19 e \nbsp 'then' \nbsp term 19 t \nbsp 'else' \nbsp term 90 f \nbsp" non associative with precedence 19 for @{ 'if_then_else $e $t $f }.
+interpretation "Formula if_then_else" 'if_then_else e t f = (if_then_else ? e t f).
+
+notation > "hvbox(a break ≤ b)" non associative with precedence 45 for @{oa_leq $a $b}.
+notation > "a >< b" non associative with precedence 45 for @{oa_overlap $a $b}.
+notation > "⋁ p" non associative with precedence 45 for @{oa_join ? $p}.
+notation > "⋀ p" non associative with precedence 45 for @{oa_meet ? $p}.
+notation > "𝟙" non associative with precedence 90 for @{oa_one}.
+notation > "𝟘" non associative with precedence 90 for @{oa_zero}.
+record OAlgebra : Type[2] := {
+ oa_P :> SET1;
+ oa_leq : oa_P × oa_P ⇒_1 CPROP;
+ oa_overlap: oa_P × oa_P ⇒_1 CPROP;
+ oa_meet: ∀I:SET.(I ⇒_2 oa_P) ⇒_2. oa_P;
+ oa_join: ∀I:SET.(I ⇒_2 oa_P) ⇒_2. oa_P;
+ oa_one: oa_P;
+ oa_zero: oa_P;
+ oa_leq_refl: ∀a:oa_P. a ≤ a;
+ oa_leq_antisym: ∀a,b:oa_P.a ≤ b → b ≤ a → a = b;
+ oa_leq_trans: ∀a,b,c:oa_P.a ≤ b → b ≤ c → a ≤ c;
+ oa_overlap_sym: ∀a,b:oa_P.a >< b → b >< a;
+ oa_meet_inf: ∀I:SET.∀p_i:I ⇒_2 oa_P.∀p:oa_P.p ≤ (⋀ p_i) = (∀i:I.p ≤ (p_i i));
+ oa_join_sup: ∀I:SET.∀p_i:I ⇒_2 oa_P.∀p:oa_P.(⋁ p_i) ≤ p = (∀i:I.p_i i ≤ p);
+ 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 | 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
+ 2) manca (vedere altro capitolo) la "suriettivita'" come immagine di insiemi di oa_base
+ oa_enum : ums oa_base oa_P;
+ oa_density: ∀p,q.(∀i.oa_overlap p (oa_enum i) → oa_overlap q (oa_enum i)) → oa_leq p q
+ *)
+ oa_density: ∀p,q.(∀r.p >< r → q >< r) → p ≤ q
+}.
+
+notation "hvbox(a break ≤ b)" non associative with precedence 45 for @{ 'leq $a $b }.
+
+interpretation "o-algebra leq" 'leq a b = (fun21 ??? (oa_leq ?) a b).
+
+notation "hovbox(a mpadded width -150% (>)< b)" non associative with precedence 45
+for @{ 'overlap $a $b}.
+interpretation "o-algebra overlap" 'overlap a b = (fun21 ??? (oa_overlap ?) a b).
+
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∧) \below (\emsp) \nbsp term 90 p)"
+non associative with precedence 50 for @{ 'oa_meet $p }.
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∧) \below (ident i ∈ I) break term 90 p)"
+non associative with precedence 50 for @{ 'oa_meet_mk (λ${ident i}:$I.$p) }.
+
+notation > "hovbox(∧ f)" non associative with precedence 60
+for @{ 'oa_meet $f }.
+interpretation "o-algebra meet" 'oa_meet f =
+ (fun12 ?? (oa_meet ??) f).
+interpretation "o-algebra meet with explicit function" 'oa_meet_mk f =
+ (fun12 ?? (oa_meet ??) (mk_unary_morphism1 ?? f ?)).
+
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∨) \below (\emsp) \nbsp term 90 p)"
+non associative with precedence 50 for @{ 'oa_join $p }.
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∨) \below (ident i ∈ I) break term 90 p)"
+non associative with precedence 50 for @{ 'oa_join_mk (λ${ident i}:$I.$p) }.
+
+notation > "hovbox(∨ f)" non associative with precedence 60
+for @{ 'oa_join $f }.
+interpretation "o-algebra join" 'oa_join f =
+ (fun12 ?? (oa_join ??) f).
+interpretation "o-algebra join with explicit function" 'oa_join_mk f =
+ (fun12 ?? (oa_join ??) (mk_unary_morphism ?? f ?)).
+
+definition binary_meet : ∀O:OAlgebra. O × O ⇒_1 O.
+intros; split;
+[ intros (p q);
+ apply (∧ { x ∈ BOOL | match x with [ true ⇒ p | false ⇒ q ] | IF_THEN_ELSE_p ? p q });
+| intros; lapply (prop12 ? O (oa_meet O BOOL));
+ [2: apply ({ x ∈ BOOL | match x with [ true ⇒ a | false ⇒ b ] | IF_THEN_ELSE_p ? a b });
+ |3: apply ({ x ∈ BOOL | match x with [ true ⇒ a' | false ⇒ b' ] | IF_THEN_ELSE_p ? a' b' });
+ | apply Hletin;]
+ intro x; simplify; cases x; simplify; assumption;]
+qed.
+
+interpretation "o-algebra binary meet" 'and a b =
+ (fun21 ??? (binary_meet ?) a b).
+
+prefer coercion Type[1]_OF_OAlgebra.
+
+definition binary_join : ∀O:OAlgebra. O × O ⇒_1 O.
+intros; split;
+[ intros (p q);
+ apply (∨ { x ∈ BOOL | match x with [ true ⇒ p | false ⇒ q ] | IF_THEN_ELSE_p ? p q });
+| intros; lapply (prop12 ? O (oa_join O BOOL));
+ [2: apply ({ x ∈ BOOL | match x with [ true ⇒ a | false ⇒ b ] | IF_THEN_ELSE_p ? a b });
+ |3: apply ({ x ∈ BOOL | match x with [ true ⇒ a' | false ⇒ b' ] | IF_THEN_ELSE_p ? a' b' });
+ | apply Hletin;]
+ intro x; simplify; cases x; simplify; assumption;]
+qed.
+
+interpretation "o-algebra binary join" 'or a b =
+ (fun21 ??? (binary_join ?) a b).
+
+lemma oa_overlap_preservers_meet: ∀O:OAlgebra.∀p,q:O.p >< q → p >< (p ∧ q).
+intros; lapply (oa_overlap_preserves_meet_ O p q f) as H; clear f;
+(** screenshot "screenoa". *)
+assumption;
+qed.
+
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∨) \below (\emsp) \nbsp term 90 p)"
+non associative with precedence 49 for @{ 'oa_join $p }.
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∨) \below (ident i ∈ I) break term 90 p)"
+non associative with precedence 49 for @{ 'oa_join_mk (λ${ident i}:$I.$p) }.
+notation < "hovbox(a ∨ b)" left associative with precedence 49
+for @{ 'oa_join_mk (λ${ident i}:$_.match $i with [ true ⇒ $a | false ⇒ $b ]) }.
+
+notation > "hovbox(∨ f)" non associative with precedence 59
+for @{ 'oa_join $f }.
+notation > "hovbox(a ∨ b)" left associative with precedence 49
+for @{ 'oa_join (mk_unary_morphism BOOL ? (λx__:bool.match x__ with [ true ⇒ $a | false ⇒ $b ]) (IF_THEN_ELSE_p ? $a $b)) }.
+
+interpretation "o-algebra join" 'oa_join f =
+ (fun12 ?? (oa_join ??) f).
+interpretation "o-algebra join with explicit function" 'oa_join_mk f =
+ (fun12 ?? (oa_join ??) (mk_unary_morphism ?? f ?)).
+
+record ORelation (P,Q : OAlgebra) : Type[2] ≝ {
+ or_f_ : P ⇒_2 Q;
+ or_f_minus_star_ : P ⇒_2 Q;
+ or_f_star_ : Q ⇒_2 P;
+ or_f_minus_ : Q ⇒_2 P;
+ or_prop1_ : ∀p,q. (or_f_ p ≤ q) = (p ≤ or_f_star_ q);
+ or_prop2_ : ∀p,q. (or_f_minus_ p ≤ q) = (p ≤ or_f_minus_star_ q);
+ or_prop3_ : ∀p,q. (or_f_ p >< q) = (p >< or_f_minus_ q)
+}.
+
+definition ORelation_setoid : OAlgebra → OAlgebra → setoid2.
+intros (P Q);
+constructor 1;
+[ apply (ORelation P Q);
+| constructor 1;
+ (* tenere solo una uguaglianza e usare la proposizione 9.9 per
+ le altre (unicita' degli aggiunti e del simmetrico) *)
+ [ apply (λp,q. And42
+ (or_f_minus_star_ ?? p = or_f_minus_star_ ?? q)
+ (or_f_minus_ ?? p = or_f_minus_ ?? q)
+ (or_f_ ?? p = or_f_ ?? q)
+ (or_f_star_ ?? p = or_f_star_ ?? q));
+ | whd; simplify; intros; repeat split; intros; apply refl2;
+ | whd; simplify; intros; cases a; clear a; split;
+ intro a; apply sym1; generalize in match a;assumption;
+ | whd; simplify; intros; cases a; cases a1; clear a a1; split; intro a;
+ [ apply (.= (e a)); apply e4;
+ | apply (.= (e1 a)); apply e5;
+ | apply (.= (e2 a)); apply e6;
+ | apply (.= (e3 a)); apply e7;]]]
+qed.
+
+definition ORelation_of_ORelation_setoid :
+ ∀P,Q.ORelation_setoid P Q → ORelation P Q ≝ λP,Q,x.x.
+coercion ORelation_of_ORelation_setoid.
+
+definition or_f_minus_star: ∀P,Q:OAlgebra.(ORelation_setoid P Q) ⇒_2 (P ⇒_2 Q).
+ intros; constructor 1;
+ [ apply or_f_minus_star_;
+ | intros; cases e; assumption]
+qed.
+
+definition or_f: ∀P,Q:OAlgebra.(ORelation_setoid P Q) ⇒_2 (P ⇒_2 Q).
+ intros; constructor 1;
+ [ apply or_f_;
+ | intros; cases e; assumption]
+qed.
+
+definition or_f_minus: ∀P,Q:OAlgebra.(ORelation_setoid P Q) ⇒_2 (Q ⇒_2 P).
+ intros; constructor 1;
+ [ apply or_f_minus_;
+ | intros; cases e; assumption]
+qed.
+
+definition or_f_star: ∀P,Q:OAlgebra.(ORelation_setoid P Q) ⇒_2 (Q ⇒_2 P).
+ intros; constructor 1;
+ [ apply or_f_star_;
+ | intros; cases e; assumption]
+qed.
+
+lemma arrows1_of_ORelation_setoid : ∀P,Q. ORelation_setoid P Q → (P ⇒_2 Q).
+intros; apply (or_f ?? c);
+qed.
+coercion arrows1_of_ORelation_setoid.
+
+interpretation "o-relation f⎻*" 'OR_f_minus_star r = (fun12 ?? (or_f_minus_star ? ?) r).
+interpretation "o-relation f⎻" 'OR_f_minus r = (fun12 ?? (or_f_minus ? ?) r).
+interpretation "o-relation f*" 'OR_f_star r = (fun12 ?? (or_f_star ? ?) r).
+
+definition or_prop1 : ∀P,Q:OAlgebra.∀F:ORelation_setoid P Q.∀p,q.
+ (F p ≤ q) =_1 (p ≤ F* q).
+intros; apply (or_prop1_ ?? F p q);
+qed.
+
+definition or_prop2 : ∀P,Q:OAlgebra.∀F:ORelation_setoid P Q.∀p,q.
+ (F⎻ p ≤ q) = (p ≤ F⎻* q).
+intros; apply (or_prop2_ ?? F p q);
+qed.
+
+definition or_prop3 : ∀P,Q:OAlgebra.∀F:ORelation_setoid P Q.∀p,q.
+ (F p >< q) = (p >< F⎻ q).
+intros; apply (or_prop3_ ?? F p q);
+qed.
+
+definition ORelation_composition : ∀P,Q,R.
+ (ORelation_setoid P Q) × (ORelation_setoid Q R) ⇒_2 (ORelation_setoid P R).
+intros;
+constructor 1;
+[ intros (F G);
+ constructor 1;
+ [ apply (G ∘ F);
+ | apply rule (G⎻* ∘ F⎻* );
+ | apply (F* ∘ G* );
+ | apply (F⎻ ∘ G⎻);
+ | intros;
+ change with ((G (F p) ≤ q) = (p ≤ (F* (G* q))));
+ apply (.= (or_prop1 :?));
+ apply (or_prop1 :?);
+ | intros;
+ change with ((F⎻ (G⎻ p) ≤ q) = (p ≤ (G⎻* (F⎻* q))));
+ apply (.= (or_prop2 :?));
+ apply or_prop2 ;
+ | intros; change with ((G (F (p)) >< q) = (p >< (F⎻ (G⎻ q))));
+ apply (.= (or_prop3 :?));
+ apply or_prop3;
+ ]
+| intros; split; simplify;
+ [3: unfold arrows1_of_ORelation_setoid; apply ((†e)‡(†e1));
+ |1: apply ((†e)‡(†e1));
+ |2,4: apply ((†e1)‡(†e));]]
+qed.
+
+definition OA : category2.
+split;
+[ apply (OAlgebra);
+| intros; apply (ORelation_setoid o o1);
+| intro O; split;
+ [1,2,3,4: apply id2;
+ |5,6,7:intros; apply refl1;]
+| apply ORelation_composition;
+| intros (P Q R S F G H); split;
+ [ change with (H⎻* ∘ G⎻* ∘ F⎻* = H⎻* ∘ (G⎻* ∘ F⎻* ));
+ apply (comp_assoc2 ????? (F⎻* ) (G⎻* ) (H⎻* ));
+ | apply ((comp_assoc2 ????? (H⎻) (G⎻) (F⎻))^-1);
+ | apply ((comp_assoc2 ????? F G H)^-1);
+ | apply ((comp_assoc2 ????? H* G* F* ));]
+| intros; split; unfold ORelation_composition; simplify; apply id_neutral_left2;
+| intros; split; unfold ORelation_composition; simplify; apply id_neutral_right2;]
+qed.
+
+definition OAlgebra_of_objs2_OA: objs2 OA → OAlgebra ≝ λx.x.
+coercion OAlgebra_of_objs2_OA.
+
+definition ORelation_setoid_of_arrows2_OA:
+ ∀P,Q. arrows2 OA P Q → ORelation_setoid P Q ≝ λP,Q,c.c.
+coercion ORelation_setoid_of_arrows2_OA.
+
+prefer coercion Type_OF_objs2.
+
+notation > "B ⇒_\o2 C" right associative with precedence 72 for @{'arrows2_OA $B $C}.
+notation "B ⇒\sub (\o 2) C" right associative with precedence 72 for @{'arrows2_OA $B $C}.
+interpretation "'arrows2_OA" 'arrows2_OA A B = (arrows2 OA A B).
+
+(* qui la notazione non va *)
+lemma leq_to_eq_join: ∀S:OA.∀p,q:S. p ≤ q → q = (binary_join ? p q).
+ intros;
+ apply oa_leq_antisym;
+ [ apply oa_density; intros;
+ apply oa_overlap_sym;
+ unfold binary_join; simplify;
+ apply (. (oa_join_split : ?));
+ exists; [ apply false ]
+ apply oa_overlap_sym;
+ assumption
+ | unfold binary_join; simplify;
+ apply (. (oa_join_sup : ?)); intro;
+ cases i; whd in ⊢ (? ? ? ? ? % ?);
+ [ assumption | apply oa_leq_refl ]]
+qed.
+
+lemma overlap_monotone_left: ∀S:OA.∀p,q,r:S. p ≤ q → p >< r → q >< r.
+ intros;
+ apply (. (leq_to_eq_join : ?)‡#);
+ [ apply f;
+ | skip
+ | apply oa_overlap_sym;
+ unfold binary_join; simplify;
+ apply (. (oa_join_split : ?));
+ exists [ apply true ]
+ apply oa_overlap_sym;
+ assumption; ]
+qed.
+
+(* Part of proposition 9.9 *)
+lemma f_minus_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R⎻ p ≤ R⎻ q.
+ intros;
+ apply (. (or_prop2 : ?));
+ apply oa_leq_trans; [2: apply f; | skip | apply (. (or_prop2 : ?)^ -1); apply oa_leq_refl;]
+qed.
+
+(* Part of proposition 9.9 *)
+lemma f_minus_star_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R⎻* p ≤ R⎻* q.
+ intros;
+ apply (. (or_prop2 : ?)^ -1);
+ apply oa_leq_trans; [3: apply f; | skip | apply (. (or_prop2 : ?)); apply oa_leq_refl;]
+qed.
+
+(* Part of proposition 9.9 *)
+lemma f_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R p ≤ R q.
+ intros;
+ apply (. (or_prop1 : ?));
+ apply oa_leq_trans; [2: apply f; | skip | apply (. (or_prop1 : ?)^ -1); apply oa_leq_refl;]
+qed.
+
+(* Part of proposition 9.9 *)
+lemma f_star_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R* p ≤ R* q.
+ intros;
+ apply (. (or_prop1 : ?)^ -1);
+ apply oa_leq_trans; [3: apply f; | skip | apply (. (or_prop1 : ?)); apply oa_leq_refl;]
+qed.
+
+lemma lemma_10_2_a: ∀S,T.∀R:arrows2 OA S T.∀p. p ≤ R⎻* (R⎻ p).
+ intros;
+ apply (. (or_prop2 : ?)^-1);
+ apply oa_leq_refl.
+qed.
+
+lemma lemma_10_2_b: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻ (R⎻* p) ≤ p.
+ intros;
+ apply (. (or_prop2 : ?));
+ apply oa_leq_refl.
+qed.
+
+lemma lemma_10_2_c: ∀S,T.∀R:arrows2 OA S T.∀p. p ≤ R* (R p).
+ intros;
+ apply (. (or_prop1 : ?)^-1);
+ apply oa_leq_refl.
+qed.
+
+lemma lemma_10_2_d: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* p) ≤ p.
+ intros;
+ apply (. (or_prop1 : ?));
+ apply oa_leq_refl.
+qed.
+
+lemma lemma_10_3_a: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻ (R⎻* (R⎻ p)) = R⎻ p.
+ intros; apply oa_leq_antisym;
+ [ apply lemma_10_2_b;
+ | apply f_minus_image_monotone;
+ apply lemma_10_2_a; ]
+qed.
+
+lemma lemma_10_3_b: ∀S,T.∀R:arrows2 OA S T.∀p. R* (R (R* p)) = R* p.
+ intros; apply oa_leq_antisym;
+ [ apply f_star_image_monotone;
+ apply (lemma_10_2_d ?? R p);
+ | apply lemma_10_2_c; ]
+qed.
+
+lemma lemma_10_3_c: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* (R p)) = R p.
+ intros; apply oa_leq_antisym;
+ [ apply lemma_10_2_d;
+ | apply f_image_monotone;
+ apply (lemma_10_2_c ?? R p); ]
+qed.
+
+lemma lemma_10_3_d: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻* (R⎻ (R⎻* p)) = R⎻* p.
+ intros; apply oa_leq_antisym;
+ [ apply f_minus_star_image_monotone;
+ apply (lemma_10_2_b ?? R p);
+ | apply lemma_10_2_a; ]
+qed.
+
+lemma lemma_10_4_a: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻* (R⎻ (R⎻* (R⎻ p))) = R⎻* (R⎻ p).
+ intros; apply (†(lemma_10_3_a ?? R p));
+qed.
+
+lemma lemma_10_4_b: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* (R (R* p))) = R (R* p).
+intros; unfold in ⊢ (? ? ? % %); apply (†(lemma_10_3_b ?? R p));
+qed.
+
+lemma oa_overlap_sym': ∀o:OA.∀U,V:o. (U >< V) = (V >< U).
+ intros; split; intro; apply oa_overlap_sym; assumption.
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/o-algebra.ma".
+include "formal_topology/notation.ma".
+
+record Obasic_pair: Type[2] ≝ {
+ Oconcr: OA; Oform: OA; Orel: arrows2 ? Oconcr Oform
+}.
+
+(* FIX *)
+interpretation "o-basic pair relation indexed" 'Vdash2 x y c = (Orel c x y).
+interpretation "o-basic pair relation (non applied)" 'Vdash c = (Orel c).
+
+record Orelation_pair (BP1,BP2: Obasic_pair): Type[2] ≝ {
+ Oconcr_rel: (Oconcr BP1) ⇒_\o2 (Oconcr BP2); Oform_rel: (Oform BP1) ⇒_\o2 (Oform BP2);
+ Ocommute: ⊩ ∘ Oconcr_rel =_2 Oform_rel ∘ ⊩
+}.
+
+(* FIX *)
+interpretation "o-concrete relation" 'concr_rel r = (Oconcr_rel ?? r).
+interpretation "o-formal relation" 'form_rel r = (Oform_rel ?? r).
+
+definition Orelation_pair_equality:
+ ∀o1,o2. equivalence_relation2 (Orelation_pair o1 o2).
+ intros;
+ constructor 1;
+ [ apply (λr,r'. ⊩ ∘ r \sub\c = ⊩ ∘ r' \sub\c);
+ | simplify;
+ intros;
+ apply refl2;
+ | simplify;
+ intros 2;
+ apply sym2;
+ | simplify;
+ intros 3;
+ apply trans2;
+ ]
+qed.
+
+(* qui setoid1 e' giusto: ma non lo e'!!! *)
+definition Orelation_pair_setoid: Obasic_pair → Obasic_pair → setoid2.
+ intros;
+ constructor 1;
+ [ apply (Orelation_pair o o1)
+ | apply Orelation_pair_equality
+ ]
+qed.
+
+definition Orelation_pair_of_Orelation_pair_setoid:
+ ∀P,Q. Orelation_pair_setoid P Q → Orelation_pair P Q ≝ λP,Q,x.x.
+coercion Orelation_pair_of_Orelation_pair_setoid.
+
+lemma eq_to_eq': ∀o1,o2.∀r,r': Orelation_pair_setoid o1 o2. r =_2 r' → r \sub\f ∘ ⊩ =_2 r'\sub\f ∘ ⊩.
+ intros 5 (o1 o2 r r' H); change in H with (⊩ ∘ r\sub\c = ⊩ ∘ r'\sub\c);
+ apply (.= ((Ocommute ?? r) ^ -1));
+ apply (.= H);
+ apply (.= (Ocommute ?? r'));
+ apply refl2;
+qed.
+
+
+definition Oid_relation_pair: ∀o:Obasic_pair. Orelation_pair o o.
+ intro;
+ constructor 1;
+ [1,2: apply id2;
+ | lapply (id_neutral_right2 ? (Oconcr o) ? (⊩)) as H;
+ lapply (id_neutral_left2 ?? (Oform o) (⊩)) as H1;
+ apply (.= H);
+ apply (H1^-1);]
+qed.
+
+lemma Orelation_pair_composition:
+ ∀o1,o2,o3:Obasic_pair.
+ Orelation_pair_setoid o1 o2 → Orelation_pair_setoid o2 o3→Orelation_pair_setoid o1 o3.
+intros 3 (o1 o2 o3);
+ intros (r r1);
+ constructor 1;
+ [ apply (r1 \sub\c ∘ r \sub\c)
+ | apply (r1 \sub\f ∘ r \sub\f)
+ | lapply (Ocommute ?? r) as H;
+ lapply (Ocommute ?? r1) as H1;
+ apply rule (.= ASSOC);
+ apply (.= #‡H1);
+ apply rule (.= ASSOC ^ -1);
+ apply (.= H‡#);
+ apply rule ASSOC]
+qed.
+
+
+lemma Orelation_pair_composition_is_morphism:
+ ∀o1,o2,o3:Obasic_pair.
+ Πa,a':Orelation_pair_setoid o1 o2.Πb,b':Orelation_pair_setoid o2 o3.
+ a=a' →b=b' →
+ Orelation_pair_composition o1 o2 o3 a b
+ = Orelation_pair_composition o1 o2 o3 a' b'.
+intros;
+ change with (⊩ ∘ (b\sub\c ∘ a\sub\c) = ⊩ ∘ (b'\sub\c ∘ a'\sub\c));
+ change in e with (⊩ ∘ a \sub\c = ⊩ ∘ a' \sub\c);
+ change in e1 with (⊩ ∘ b \sub\c = ⊩ ∘ b' \sub\c);
+ apply rule (.= ASSOC);
+ apply (.= #‡e1);
+ apply (.= #‡(Ocommute ?? b'));
+ apply rule (.= ASSOC^-1);
+ apply (.= e‡#);
+ apply rule (.= ASSOC);
+ apply (.= #‡(Ocommute ?? b')^-1);
+ apply rule (ASSOC^-1);
+qed.
+
+definition Orelation_pair_composition_morphism:
+ ∀o1,o2,o3. binary_morphism2 (Orelation_pair_setoid o1 o2) (Orelation_pair_setoid o2 o3) (Orelation_pair_setoid o1 o3).
+intros; constructor 1;
+[ apply Orelation_pair_composition;
+| apply Orelation_pair_composition_is_morphism;]
+qed.
+
+lemma Orelation_pair_composition_morphism_assoc:
+∀o1,o2,o3,o4:Obasic_pair
+ .Πa12:Orelation_pair_setoid o1 o2
+ .Πa23:Orelation_pair_setoid o2 o3
+ .Πa34:Orelation_pair_setoid o3 o4
+ .Orelation_pair_composition_morphism o1 o3 o4
+ (Orelation_pair_composition_morphism o1 o2 o3 a12 a23) a34
+ =Orelation_pair_composition_morphism o1 o2 o4 a12
+ (Orelation_pair_composition_morphism o2 o3 o4 a23 a34).
+ intros;
+ change with (⊩ ∘ (a34\sub\c ∘ (a23\sub\c ∘ a12\sub\c)) =
+ ⊩ ∘ ((a34\sub\c ∘ a23\sub\c) ∘ a12\sub\c));
+ apply rule (ASSOC‡#);
+qed.
+
+lemma Orelation_pair_composition_morphism_respects_id:
+Πo1:Obasic_pair
+.Πo2:Obasic_pair
+ .Πa:Orelation_pair_setoid o1 o2
+ .Orelation_pair_composition_morphism o1 o1 o2 (Oid_relation_pair o1) a=a.
+ intros;
+ change with (⊩ ∘ (a\sub\c ∘ (Oid_relation_pair o1)\sub\c) = ⊩ ∘ a\sub\c);
+ apply ((id_neutral_right2 ????)‡#);
+qed.
+
+lemma Orelation_pair_composition_morphism_respects_id_r:
+Πo1:Obasic_pair
+.Πo2:Obasic_pair
+ .Πa:Orelation_pair_setoid o1 o2
+ .Orelation_pair_composition_morphism o1 o2 o2 a (Oid_relation_pair o2)=a.
+intros;
+ change with (⊩ ∘ ((Oid_relation_pair o2)\sub\c ∘ a\sub\c) = ⊩ ∘ a\sub\c);
+ apply ((id_neutral_left2 ????)‡#);
+qed.
+
+definition OBP: category2.
+ constructor 1;
+ [ apply Obasic_pair
+ | apply Orelation_pair_setoid
+ | apply Oid_relation_pair
+ | apply Orelation_pair_composition_morphism
+ | apply Orelation_pair_composition_morphism_assoc;
+ | apply Orelation_pair_composition_morphism_respects_id;
+ | apply Orelation_pair_composition_morphism_respects_id_r;]
+qed.
+
+definition Obasic_pair_of_objs2_OBP: objs2 OBP → Obasic_pair ≝ λx.x.
+coercion Obasic_pair_of_objs2_OBP.
+
+definition Orelation_pair_setoid_of_arrows2_OBP:
+ ∀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;
+ [ apply (ext ? ? (rel o));
+ | intros;
+ apply (.= #‡H);
+ apply refl1]
+qed.
+
+definition BPextS: ∀o: BP. Ω \sup (form o) ⇒ Ω \sup (concr o) ≝
+ λo.extS ?? (rel o).
+*)
+
+(*
+definition fintersects: ∀o: BP. binary_morphism1 (form o) (form o) (Ω \sup (form o)).
+ intros (o); constructor 1;
+ [ apply (λa,b: form o.{c | BPext o c ⊆ BPext o a ∩ BPext o b });
+ intros; simplify; apply (.= (†H)‡#); apply refl1
+ | intros; split; simplify; intros;
+ [ apply (. #‡((†H)‡(†H1))); assumption
+ | apply (. #‡((†H\sup -1)‡(†H1\sup -1))); assumption]]
+qed.
+
+interpretation "fintersects" 'fintersects U V = (fun1 ??? (fintersects ?) U V).
+
+definition fintersectsS:
+ ∀o:BP. binary_morphism1 (Ω \sup (form o)) (Ω \sup (form o)) (Ω \sup (form o)).
+ intros (o); constructor 1;
+ [ apply (λo: basic_pair.λa,b: Ω \sup (form o).{c | BPext o c ⊆ BPextS o a ∩ BPextS o b });
+ intros; simplify; apply (.= (†H)‡#); apply refl1
+ | intros; split; simplify; intros;
+ [ apply (. #‡((†H)‡(†H1))); assumption
+ | apply (. #‡((†H\sup -1)‡(†H1\sup -1))); assumption]]
+qed.
+
+interpretation "fintersectsS" 'fintersects U V = (fun1 ??? (fintersectsS ?) U V).
+*)
+
+(*
+definition relS: ∀o: BP. binary_morphism1 (concr o) (Ω \sup (form o)) CPROP.
+ intros (o); constructor 1;
+ [ apply (λx:concr o.λS: Ω \sup (form o).∃y: form o.y ∈ S ∧ x ⊩ y);
+ | intros; split; intros; cases H2; exists [1,3: apply w]
+ [ apply (. (#‡H1)‡(H‡#)); assumption
+ | apply (. (#‡H1 \sup -1)‡(H \sup -1‡#)); assumption]]
+qed.
+
+interpretation "basic pair relation for subsets" 'Vdash2 x y = (fun1 (concr ?) ?? (relS ?) x y).
+interpretation "basic pair relation for subsets (non applied)" 'Vdash = (fun1 ??? (relS ?)).
+*)
+
+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 = (fun12 ? ? (or_f_minus_star ? ?) (Orel 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 = (fun12 ? ? (or_f ? ?) (Orel 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 = (fun12 ? ? (or_f_star ? ?) (Orel 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 = (fun12 ? ? (or_f_minus ? ?) (Orel x)).
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/notation.ma".
+include "formal_topology/o-basic_pairs.ma".
+include "formal_topology/o-basic_topologies.ma".
+
+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 → OBTop.
+ intro t;
+ constructor 1;
+ [ apply (Oform t);
+ | apply (□⎽t ∘ Ext⎽t);
+ | apply (◊⎽t ∘ Rest⎽t);
+ | apply hide; intros 2; split; intro;
+ [ change with ((⊩) \sup ⎻* ((⊩) \sup ⎻ U) ≤ (⊩) \sup ⎻* ((⊩) \sup ⎻ V));
+ apply (. (#‡(lemma_10_4_a ?? (⊩) V)^-1));
+ apply f_minus_star_image_monotone;
+ apply f_minus_image_monotone;
+ assumption
+ | apply oa_leq_trans;
+ [3: apply f;
+ | skip
+ | change with (U ≤ (⊩)⎻* ((⊩)⎻ U));
+ apply (. (or_prop2 : ?) ^ -1);
+ apply oa_leq_refl; ]]
+ | apply hide; intros 2; split; intro;
+ [ change with (◊⎽t ((⊩) \sup * U) ≤ ◊⎽t ((⊩) \sup * V));
+ apply (. ((lemma_10_4_b ?? (⊩) U)^-1)‡#);
+ apply (f_image_monotone ?? (⊩) ? ((⊩)* V));
+ apply f_star_image_monotone;
+ assumption;
+ | apply oa_leq_trans;
+ [2: apply f;
+ | skip
+ | change with ((⊩) ((⊩)* V) ≤ V);
+ apply (. (or_prop1 : ?));
+ apply oa_leq_refl; ]]
+ | apply hide; intros;
+ apply (.= (oa_overlap_sym' : ?));
+ change with ((◊⎽t ((⊩)* V) >< (⊩)⎻* ((⊩)⎻ U)) = (U >< (◊⎽t ((⊩)* V))));
+ apply (.= (or_prop3 ?? (⊩) ((⊩)* V) ?));
+ apply (.= #‡(lemma_10_3_a : ?));
+ apply (.= (or_prop3 : ?)^-1);
+ apply (oa_overlap_sym' ? ((⊩) ((⊩)* V)) U); ]
+qed.
+
+definition o_continuous_relation_of_o_relation_pair:
+ ∀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);
+ | apply hide; unfold o_basic_topology_of_o_basic_pair; simplify; intros (U e);
+ apply sym1;
+ apply (.= †(†e));
+ change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f ∘ (⊩)) ((⊩\sub BP1)* U));
+ cut ((t \sub \f ∘ (⊩)) ((⊩\sub BP1)* U) = ((⊩) ∘ t \sub \c) ((⊩\sub BP1)* U)) as COM;[2:
+ cases (Ocommute ?? t); apply (e3 ^ -1 ((⊩\sub BP1)* U));]
+ apply (.= †COM);
+ change in ⊢ (? ? ? % ?) with (((⊩) ∘ (⊩)* ) (((⊩) ∘ t \sub \c ∘ (⊩)* ) U));
+ apply (.= (lemma_10_3_c ?? (⊩) (t \sub \c ((⊩\sub BP1)* U))));
+ apply (.= COM ^ -1);
+ change in ⊢ (? ? ? % ?) with (t \sub \f (((⊩) ∘ (⊩\sub BP1)* ) U));
+ change in e with (U=((⊩)∘(⊩ \sub BP1) \sup * ) U);
+ apply (†e^-1);
+ | apply hide; unfold o_basic_topology_of_o_basic_pair; simplify; intros;
+ apply sym1;
+ apply (.= †(†e));
+ change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f⎻* ∘ (⊩)⎻* ) ((⊩\sub BP1)⎻ U));
+ cut ((t \sub \f⎻* ∘ (⊩)⎻* ) ((⊩\sub BP1)⎻ U) = ((⊩)⎻* ∘ t \sub \c⎻* ) ((⊩\sub BP1)⎻ U)) as COM;[2:
+ cases (Ocommute ?? t); apply (e1 ^ -1 ((⊩\sub BP1)⎻ U));]
+ apply (.= †COM);
+ change in ⊢ (? ? ? % ?) with (((⊩)⎻* ∘ (⊩)⎻ ) (((⊩)⎻* ∘ t \sub \c⎻* ∘ (⊩)⎻ ) U));
+ apply (.= (lemma_10_3_d ?? (⊩) (t \sub \c⎻* ((⊩\sub BP1)⎻ U))));
+ apply (.= COM ^ -1);
+ change in ⊢ (? ? ? % ?) with (t \sub \f⎻* (((⊩)⎻* ∘ (⊩\sub BP1)⎻ ) U));
+ change in e with (U=((⊩)⎻* ∘(⊩ \sub BP1)⎻ ) U);
+ apply (†e^-1);]
+qed.
+
+
+definition OR : carr3 (OBP ⇒_\c3 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 ⎻* ∘ 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 (.= #‡†(Ocommute:?)^-1);
+ apply (.= #‡e1);
+ change in ⊢ (? ? ? (? ? ? ? ? ? %) ?) with (⊩\sub T ∘ a'\sub\c)⎻*;
+ apply (.= #‡†(Ocommute:?));
+ change in ⊢ (? ? ? (? ? ? ? ? ? %) ?) with (a'\sub\f⎻* ∘ (⊩\sub S)⎻* );
+ apply (.= (comp_assoc2 ? ???? ?? a'\sub\f⎻* )^-1);
+ apply refl2;]
+| intros 2 (o a); apply refl1;
+| intros 6; apply refl1;]
+qed.
+
--- /dev/null
+ (**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/o-algebra.ma".
+include "formal_topology/o-saturations.ma".
+
+record Obasic_topology: Type[2] ≝ {
+ Ocarrbt:> OA;
+ oA: Ocarrbt ⇒_2 Ocarrbt; oJ: Ocarrbt ⇒_2 Ocarrbt;
+ oA_is_saturation: is_o_saturation ? oA; oJ_is_reduction: is_o_reduction ? oJ;
+ Ocompatibility: ∀U,V. (oA U >< oJ V) =_1 (U >< oJ V)
+ }.
+
+record Ocontinuous_relation (S,T: Obasic_topology) : Type[2] ≝ {
+ Ocont_rel:> arrows2 OA S T;
+ Oreduced: ∀U:S. U = oJ ? U → Ocont_rel U =_1 oJ ? (Ocont_rel U);
+ Osaturated: ∀U:S. U = oA ? U → Ocont_rel⎻* U =_1 oA ? (Ocont_rel⎻* U)
+ }.
+
+definition Ocontinuous_relation_setoid: Obasic_topology → Obasic_topology → setoid2.
+ intros (S T); constructor 1;
+ [ apply (Ocontinuous_relation S T)
+ | constructor 1;
+ [ alias symbol "eq" = "setoid2 eq".
+ alias symbol "compose" = "category2 composition".
+ 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 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':
+ ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
+ a = a' → ∀X.a⎻* (A o1 X) = a'⎻* (A o1 X).
+ intros; apply oa_leq_antisym; intro; unfold minus_star_image; simplify; intros;
+ [ cut (ext ?? a a1 ⊆ A ? X); [2: intros 2; apply (H1 a2); cases f1; assumption;]
+ lapply (if ?? (A_is_saturation ???) Hcut); clear Hcut;
+ cut (A ? (ext ?? a' a1) ⊆ A ? X); [2: apply (. (H ?)‡#); assumption]
+ lapply (fi ?? (A_is_saturation ???) Hcut);
+ apply (Hletin1 x); change with (x ∈ ext ?? a' a1); split; simplify;
+ [ apply I | assumption ]
+ | cut (ext ?? a' a1 ⊆ A ? X); [2: intros 2; apply (H1 a2); cases f1; assumption;]
+ lapply (if ?? (A_is_saturation ???) Hcut); clear Hcut;
+ cut (A ? (ext ?? a a1) ⊆ A ? X); [2: apply (. (H ?)\sup -1‡#); assumption]
+ lapply (fi ?? (A_is_saturation ???) Hcut);
+ apply (Hletin1 x); change with (x ∈ ext ?? a a1); split; simplify;
+ [ apply I | assumption ]]
+qed.
+
+theorem continuous_relation_eq_inv':
+ ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
+ (∀X.a⎻* (A o1 X) = a'⎻* (A o1 X)) → a=a'.
+ intros 6;
+ cut (∀a,a': continuous_relation_setoid o1 o2.
+ (∀X.a⎻* (A o1 X) = a'⎻* (A o1 X)) →
+ ∀V:(oa_P (carrbt o2)). A o1 (a'⎻ V) ≤ A o1 (a⎻ V));
+ [2: clear b H a' a; intros;
+ lapply depth=0 (λV.saturation_expansive ??? (extS ?? a V)); [2: apply A_is_saturation;|skip]
+ (* fundamental adjunction here! to be taken out *)
+ cut (∀V:Ω \sup o2.V ⊆ minus_star_image ?? a (A ? (extS ?? a V)));
+ [2: intro; intros 2; unfold minus_star_image; simplify; intros;
+ apply (Hletin V1 x); whd; split; [ exact I | exists; [apply a1] split; assumption]]
+ clear Hletin;
+ cut (∀V:Ω \sup o2.V ⊆ minus_star_image ?? a' (A ? (extS ?? a V)));
+ [2: intro; apply (. #‡(H ?)); apply Hcut] clear H Hcut;
+ (* second half of the fundamental adjunction here! to be taken out too *)
+ intro; lapply (Hcut1 (singleton ? V)); clear Hcut1;
+ unfold minus_star_image in Hletin; unfold singleton in Hletin; simplify in Hletin;
+ whd in Hletin; whd in Hletin:(?→?→%); simplify in Hletin;
+ apply (if ?? (A_is_saturation ???));
+ intros 2 (x H); lapply (Hletin V ? x ?);
+ [ apply refl | cases H; assumption; ]
+ change with (x ∈ A ? (ext ?? a V));
+ apply (. #‡(†(extS_singleton ????)));
+ assumption;]
+ split; apply Hcut; [2: assumption | intro; apply sym1; apply H]
+qed.
+*)
+
+definition Ocontinuous_relation_comp:
+ ∀o1,o2,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 (.= (Oreduced : ?)^-1);
+ [ apply (.= (Oreduced :?)); [ assumption | apply refl1 ]
+ | apply refl1]
+ | intros;
+ apply sym1;
+ change in match ((s ∘ r)⎻* U) with (s⎻* (r⎻* U));
+ apply (.= (Osaturated : ?)^-1);
+ [ apply (.= (Osaturated : ?)); [ assumption | apply refl1 ]
+ | apply refl1]]
+qed.
+
+definition OBTop: category2.
+ constructor 1;
+ [ apply Obasic_topology
+ | apply Ocontinuous_relation_setoid
+ | intro; constructor 1;
+ [ apply id2
+ | intros; apply e;
+ | intros; apply e;]
+ | intros; constructor 1;
+ [ apply Ocontinuous_relation_comp;
+ | intros; simplify;
+ 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 (b⎻* (a'⎻* (oA o1 x)) =_1 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 (b'⎻* (oA o2 (a'⎻* (oA o1 x))) =_1 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⎻* ) ∘ oA o1 = ((a34⎻* ∘ (a23⎻* ∘ a12⎻* )) ∘ oA o1));
+ apply rule (#‡ASSOC ^ -1);
+ | intros; simplify;
+ change with ((a⎻* ∘ (id2 ? o1)⎻* ) ∘ oA o1 = a⎻* ∘ oA o1);
+ apply (#‡(id_neutral_right2 : ?));
+ | intros; simplify;
+ change with (((id2 ? o2)⎻* ∘ a⎻* ) ∘ oA o1 = a⎻* ∘ oA o1);
+ apply (#‡(id_neutral_left2 : ?));]
+qed.
+
+definition Obasic_topology_of_OBTop: objs2 OBTop → Obasic_topology ≝ λx.x.
+coercion Obasic_topology_of_OBTop.
+
+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.
+
+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 *)
+theorem continuous_relation_eqS:
+ ∀o1,o2:basic_topology.∀a,a': continuous_relation_setoid o1 o2.
+ a = a' → ∀X. A ? (extS ?? a X) = A ? (extS ?? a' X).
+ intros;
+ cut (∀a: arrows1 ? o1 ?.∀x. x ∈ extS ?? a X → ∃y:o2.y ∈ X ∧ x ∈ ext ?? a y);
+ [2: intros; cases f; clear f; cases H1; exists [apply w] cases x1; split;
+ try assumption; split; assumption]
+ cut (∀a,a':continuous_relation_setoid o1 o2.eq1 ? a a' → ∀x. x ∈ extS ?? a X → ∃y:o2. y ∈ X ∧ x ∈ A ? (ext ?? a' y));
+ [2: intros; cases (Hcut ?? f); exists; [apply w] cases x1; split; try assumption;
+ apply (. #‡(H1 ?));
+ apply (saturation_expansive ?? (A_is_saturation o1) (ext ?? a1 w) x);
+ assumption;] clear Hcut;
+ split; apply (if ?? (A_is_saturation ???)); intros 2;
+ [lapply (Hcut1 a a' H a1 f) | lapply (Hcut1 a' a (H \sup -1) a1 f)]
+ cases Hletin; clear Hletin; cases x; clear x;
+ cut (∀a: arrows1 ? o1 ?. ext ?? a w ⊆ extS ?? a X);
+ [2,4: intros 3; cases f3; clear f3; simplify in f5; split; try assumption;
+ exists [1,3: apply w] split; assumption;]
+ cut (∀a. A ? (ext o1 o2 a w) ⊆ A ? (extS o1 o2 a X));
+ [2,4: intros; apply saturation_monotone; try (apply A_is_saturation); apply Hcut;]
+ apply Hcut2; assumption.
+qed.
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/o-basic_pairs.ma".
+include "formal_topology/o-saturations.ma".
+
+definition A : ∀b:OBP. unary_morphism1 (Oform b) (Oform b).
+intros; constructor 1;
+ [ apply (λx.□⎽b (Ext⎽b x));
+ | intros; apply (†(†e));]
+qed.
+
+lemma down_p : ∀S:SET1.∀I:SET.∀u:S ⇒_1 S.∀c:arrows2 SET1 I S.∀a:I.∀a':I.a =_1 a'→u (c a) =_1 u (c a').
+intros; apply (†(†e));
+qed.
+
+record Oconcrete_space : Type[2] ≝
+ { Obp:> OBP;
+ (*distr : is_distributive (form bp);*)
+ Odownarrow: unary_morphism1 (Oform Obp) (Oform Obp);
+ Odownarrow_is_sat: is_o_saturation ? Odownarrow;
+ Oconverges: ∀q1,q2.
+ (Ext⎽Obp q1 ∧ (Ext⎽Obp q2)) = (Ext⎽Obp ((Odownarrow q1) ∧ (Odownarrow q2)));
+ Oall_covered: Ext⎽Obp (oa_one (Oform Obp)) = oa_one (Oconcr Obp);
+ Oil2: ∀I:SET.∀p:arrows2 SET1 I (Oform Obp).
+ Odownarrow (∨ { x ∈ I | Odownarrow (p x) | down_p ???? }) =
+ ∨ { x ∈ I | Odownarrow (p x) | down_p ???? };
+ Oil1: ∀q.Odownarrow (A ? q) = A ? q
+ }.
+
+interpretation "o-concrete space downarrow" 'downarrow x =
+ (fun11 ?? (Odownarrow ?) x).
+
+definition Obinary_downarrow :
+ ∀C:Oconcrete_space.binary_morphism1 (Oform C) (Oform C) (Oform C).
+intros; constructor 1;
+[ intros; apply (↓ c ∧ ↓ c1);
+| intros;
+ alias symbol "prop2" = "prop21".
+ alias symbol "prop1" = "prop11".
+ apply ((†e)‡(†e1));]
+qed.
+
+interpretation "concrete_space binary ↓" 'fintersects a b = (fun21 ? ? ? (Obinary_downarrow ?) a b).
+
+record Oconvergent_relation_pair (CS1,CS2: Oconcrete_space) : Type[2] ≝
+ { Orp:> arrows2 ? CS1 CS2;
+ Orespects_converges:
+ ∀b,c. eq1 ? (Orp\sub\c⎻ (Ext⎽CS2 (b ↓ c))) (Ext⎽CS1 (Orp\sub\f⎻ b ↓ Orp\sub\f⎻ c));
+ Orespects_all_covered:
+ eq1 ? (Orp\sub\c⎻ (Ext⎽CS2 (oa_one (Oform CS2))))
+ (Ext⎽CS1 (oa_one (Oform CS1)))
+ }.
+
+definition Oconvergent_relation_space_setoid: Oconcrete_space → Oconcrete_space → setoid2.
+ intros (c c1);
+ constructor 1;
+ [ apply (Oconvergent_relation_pair c c1)
+ | constructor 1;
+ [ intros (c2 c3);
+ apply (Orelation_pair_equality c c1 c2 c3);
+ | intros 1; apply refl2;
+ | intros 2; apply sym2;
+ | intros 3; apply trans2]]
+qed.
+
+definition Oconvergent_relation_space_of_Oconvergent_relation_space_setoid:
+ ∀CS1,CS2.carr2 (Oconvergent_relation_space_setoid CS1 CS2) →
+ Oconvergent_relation_pair CS1 CS2 ≝ λP,Q,c.c.
+coercion Oconvergent_relation_space_of_Oconvergent_relation_space_setoid.
+
+definition Oconvergent_relation_space_composition:
+ ∀o1,o2,o3: Oconcrete_space.
+ binary_morphism2
+ (Oconvergent_relation_space_setoid o1 o2)
+ (Oconvergent_relation_space_setoid o2 o3)
+ (Oconvergent_relation_space_setoid o1 o3).
+ intros; constructor 1;
+ [ intros; whd in t t1 ⊢ %;
+ constructor 1;
+ [ apply (c1 ∘ c);
+ | intros;
+ change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (b↓c2))));
+ alias symbol "trans" = "trans1".
+ apply (.= († (Orespects_converges : ?)));
+ apply (Orespects_converges ?? c (c1\sub\f⎻ b) (c1\sub\f⎻ c2));
+ | change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (oa_one (Oform o3)))));
+ apply (.= (†(Orespects_all_covered :?)));
+ apply rule (Orespects_all_covered ?? c);]
+ | intros;
+ change with (b ∘ a = b' ∘ a');
+ change in e with (Orp ?? a = Orp ?? a');
+ change in e1 with (Orp ?? b = Orp ?? b');
+ apply (e‡e1);]
+qed.
+
+definition OCSPA: category2.
+ constructor 1;
+ [ apply Oconcrete_space
+ | apply Oconvergent_relation_space_setoid
+ | intro; constructor 1;
+ [ apply id2
+ | intros; apply refl1;
+ | apply refl1]
+ | apply Oconvergent_relation_space_composition
+ | intros; simplify; whd in a12 a23 a34;
+ change with (a34 ∘ (a23 ∘ a12) = (a34 ∘ a23) ∘ a12);
+ apply rule ASSOC;
+ | intros; simplify;
+ change with (a ∘ id2 OBP o1 = a);
+ apply (id_neutral_right2 : ?);
+ | intros; simplify;
+ change with (id2 ? o2 ∘ a = a);
+ apply (id_neutral_left2 : ?);]
+qed.
+
+definition Oconcrete_space_of_OCSPA : objs2 OCSPA → Oconcrete_space ≝ λx.x.
+coercion Oconcrete_space_of_OCSPA.
+
+definition Oconvergent_relation_space_setoid_of_arrows2_OCSPA :
+ ∀P,Q. arrows2 OCSPA P Q → Oconvergent_relation_space_setoid P Q ≝ λP,Q,x.x.
+coercion Oconvergent_relation_space_setoid_of_arrows2_OCSPA.
+
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/o-basic_topologies.ma".
+
+(*
+(*
+definition downarrow: ∀S:BTop. unary_morphism (Ω \sup S) (Ω \sup S).
+ intros; constructor 1;
+ [ apply (λU:Ω \sup S.{a | ∃b:carrbt S. b ∈ U ∧ a ∈ A ? (singleton ? b)});
+ intros; simplify; split; intro; cases H1; cases x; exists [1,3: apply w]
+ split; try assumption; [ apply (. H‡#) | apply (. H \sup -1‡#) ] assumption
+ | intros; split; intros 2; cases f; exists; [1,3: apply w] cases x; split;
+ try assumption; [ apply (. #‡H) | apply (. #‡H \sup -1)] assumption]
+qed.
+
+interpretation "downarrow" 'downarrow a = (fun_1 ?? (downarrow ?) a).
+
+definition ffintersects: ∀S:BTop. binary_morphism1 (Ω \sup S) (Ω \sup S) (Ω \sup S).
+ intros; constructor 1;
+ [ apply (λU,V. ↓U ∩ ↓V);
+ | intros; apply (.= (†H)‡(†H1)); apply refl1]
+qed.
+
+interpretation "ffintersects" 'fintersects U V = (fun1 ??? (ffintersects ?) U V).
+*)
+
+record formal_topology: Type ≝
+ { bt:> OBTop;
+ converges: ∀U,V: bt. oA bt (U ↓ V) = (oA ? U ∧ oA ? V)
+ }.
+
+(*
+
+definition ffintersects': ∀S:BTop. binary_morphism1 S S (Ω \sup S).
+ intros; constructor 1;
+ [ apply (λb,c:S. (singleton S b) ↓ (singleton S c));
+ | intros; apply (.= (†H)‡(†H1)); apply refl1]
+qed.
+
+interpretation "ffintersects'" 'fintersects U V = (fun1 ??? (ffintersects' ?) U V).
+*)
+record formal_map (S,T: formal_topology) : Type ≝
+ { cr:> continuous_relation_setoid S T;
+ C1: ∀b,c. extS ?? cr (b ↓ c) = ext ?? cr b ↓ ext ?? cr c;
+ C2: extS ?? cr T = S
+ }.
+
+definition formal_map_setoid: formal_topology → formal_topology → setoid1.
+ intros (S T); constructor 1;
+ [ apply (formal_map S T);
+ | constructor 1;
+ [ apply (λf,f1: formal_map S T.f=f1);
+ | simplify; intros 1; apply refl1
+ | simplify; intros 2; apply sym1
+ | simplify; intros 3; apply trans1]]
+qed.
+
+axiom C1':
+ ∀S,T: formal_topology.∀f:formal_map_setoid S T.∀U,V: Ω \sup T.
+ extS ?? f (U ↓ V) = extS ?? f U ↓ extS ?? f V.
+
+definition formal_map_composition:
+ ∀o1,o2,o3: formal_topology.
+ binary_morphism1
+ (formal_map_setoid o1 o2)
+ (formal_map_setoid o2 o3)
+ (formal_map_setoid o1 o3).
+ intros; constructor 1;
+ [ intros; whd in c c1; constructor 1;
+ [ apply (comp1 BTop ??? c c1);
+ | intros;
+ apply (.= (extS_com ??? c c1 ?));
+ apply (.= †(C1 ?????));
+ apply (.= (C1' ?????));
+ apply (.= ((†((extS_singleton ????) \sup -1))‡(†((extS_singleton ????) \sup -1))));
+ apply (.= (extS_com ??????)\sup -1 ‡ (extS_com ??????) \sup -1);
+ apply (.= (extS_singleton ????)‡(extS_singleton ????));
+ apply refl1;
+ | apply (.= (extS_com ??? c c1 ?));
+ apply (.= (†(C2 ???)));
+ apply (.= (C2 ???));
+ apply refl1;]
+ | intros; simplify;
+ change with (comp1 BTop ??? a b = comp1 BTop ??? a' b');
+ apply prop1; assumption]
+qed.
+*)
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/o-algebra.ma".
+
+definition is_o_saturation: ∀C:OA. C ⇒_1 C → CProp[1] ≝
+ λC:OA.λA:C ⇒_1 C.∀U,V. (U ≤ A V) =_1 (A U ≤ A V).
+
+definition is_o_reduction: ∀C:OA. C ⇒_1 C → CProp[1] ≝
+ λC:OA.λJ:C ⇒_1 C.∀U,V. (J U ≤ V) =_1 (J U ≤ J V).
+
+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 o_saturation_monotone: ∀C:OA.∀A:C ⇒_1 C. 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 o_saturation_expansive ]
+ assumption.
+qed.
+
+theorem o_saturation_idempotent: ∀C:OA.∀A:C ⇒_1 C. is_o_saturation C A → ∀U. A (A U) =_1 A U.
+ intros; apply (oa_leq_antisym C);
+ [ apply (if ?? (i (A U) U)); apply (oa_leq_refl C).
+ | apply o_saturation_expansive; assumption]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/basic_pairs_to_o-basic_pairs.ma".
+include "formal_topology/apply_functor.ma".
+
+definition rOBP ≝ Apply (category2_of_category1 BP) OBP BP_to_OBP.
+
+include "formal_topology/o-basic_pairs_to_o-basic_topologies.ma".
+
+lemma category2_of_category1_respects_comp_r:
+ ∀C:category1.∀o1,o2,o3:C.
+ ∀f:arrows1 ? o1 o2.∀g:arrows1 ? o2 o3.
+ (comp1 ???? f g) =_\ID (comp2 (category2_of_category1 C) o1 o2 o3 f g).
+ intros; constructor 1;
+qed.
+
+lemma category2_of_category1_respects_comp:
+ ∀C:category1.∀o1,o2,o3:C.
+ ∀f:arrows1 ? o1 o2.∀g:arrows1 ? o2 o3.
+ (comp2 (category2_of_category1 C) o1 o2 o3 f g) =_\ID (comp1 ???? f g).
+ intros; constructor 1;
+qed.
+
+lemma POW_full':
+ ∀S,T:REL.∀f:arrows2 SET1 (POW S) (POW T).
+ arrows1 REL S T.
+ intros;
+ constructor 1; constructor 1;
+ [ intros (x y); apply (y ∈ c {(x)});
+ | apply hide; intros; unfold FunClass_1_OF_Ocontinuous_relation;
+ apply (e1ࠠe); ]
+qed.
+
+(*
+lemma POW_full_image:
+ ∀S,T:REL.∀f:arrows2 SET1 (POW S) (POW T).
+ exT22 ? (λg:arrows1 REL S T.or_f ?? (map_arrows2 ?? POW S T g) = f).
+ intros; letin g ≝ (? : carr1 (arrows1 REL S T)); [
+ constructor 1; constructor 1;
+ [ intros (x y); apply (y ∈ f {(x)});
+ | apply hide; intros; unfold FunClass_1_OF_Ocontinuous_relation;
+ apply (e1ࠠe); ]]
+exists [apply g]
+intro; split; intro; simplify; intro;
+[ whd in f1; change in f1:(? ? (λ_:?.? ? ? ? ? % ?)) with (a1 ∈ f {(x)});
+ cases f1; cases x; clear f1 x; change with (a1 ∈ f a);
+ lapply (f_image_monotone ?? (map_arrows2 ?? POW S T g) (singleton ? w) a ? a1);
+ [2: whd in Hletin;
+ change in Hletin:(? ? (λ_:?.? ? ? ? ? % ?))
+ with (a1 ∈ f {(x)}); cases Hletin; cases x;
+ [ intros 2; change in f3 with (eq1 ? w a2); change with (a2 ∈ a);
+ apply (. f3^-1‡#); assumption;
+ | assumption; ]
+
+
+
+ lapply (. (or_prop3 ?? (map_arrows2 ?? POW S T g) (singleton ? a1) a)^-1);
+ [ whd in Hletin:(? ? ? ? ? ? %);
+ change in Hletin:(? ? ? ? ? ? (? ? (? ? ? (λ_:?.? ? (λ_:?.? ? ? ? ? % ?)) ?)))
+ with (y ∈ f {(x)});
+ cases Hletin; cases x1; cases x2;
+
+ [ cases Hletin; change in x with (eq1 ? a1 w1); apply (. x‡#); assumption;
+ | exists; [apply w] assumption ]
+
+
+ clear g;
+ cases f1; cases x; simplify in f2; change with (a1 ∈ (f a));
+ lapply depth=0 (let x ≝ POW in or_prop3 (POW S) (POW T) (map_arrows2 ?? POW S T g));
+ lapply (Hletin {(w)} {(a1)}).
+ lapply (if ?? Hletin1); [2: clear Hletin Hletin1;
+ exists; [apply a1] [whd; exists[apply w] split; [assumption;|change with (w = w); apply rule #]]
+ change with (a1=a1); apply rule #;]
+ clear Hletin Hletin1; cases Hletin2; whd in x2;
+qed.
+*)
+lemma curry: ∀A,B,C.(A × B ⇒_1 C) → A → (B ⇒_1 C).
+ intros;
+ constructor 1;
+ [ apply (b c);
+ | intros; apply (#‡e); ]
+qed.
+
+notation < "F x" left associative with precedence 60 for @{'map_arrows $F $x}.
+interpretation "map arrows 2" 'map_arrows F x = (fun12 ? ? (map_arrows2 ? ? F ? ?) x).
+
+definition preserve_sup : ∀S,T.∀ f:Ω^S ⇒_1 Ω^T. CProp[1].
+intros (S T f); apply (∀X:Ω \sup S. (f X) =_1 ?);
+constructor 1; constructor 1;
+[ intro y; alias symbol "singl" = "singleton". alias symbol "and" = "and_morphism".
+ apply (∃x:S. x ∈ X ∧ y ∈ f {(x)});
+| intros (a b H); split; intro E; cases E; clear E; exists; [1,3:apply w]
+ [ apply (. #‡(H^-1‡#)); | apply (. #‡(H‡#));] assumption]
+qed.
+
+alias symbol "singl" = "singleton".
+lemma eq_cones_to_eq_rel:
+ ∀S,T. ∀f,g: arrows1 REL S T.
+ (∀x. curry ??? (image ??) f {(x)} = curry ??? (image ??) g {(x)}) → f = g.
+intros; intros 2 (a b); split; intro;
+[ cases (f1 a); lapply depth=0 (s b); clear s s1;
+ lapply (Hletin); clear Hletin;
+ [ cases Hletin1; cases x; change in f4 with (a = w);
+ change with (a ♮g b); apply (. f4‡#); assumption;
+ | exists; [apply a] split; [ assumption | change with (a=a); apply rule #;]]
+| cases (f1 a); lapply depth=0 (s1 b); clear s s1;
+ lapply (Hletin); clear Hletin;
+ [ cases Hletin1; cases x; change in f4 with (a = w);
+ change with (a ♮f b); apply (. f4‡#); assumption;
+ | exists; [apply a] split; [ assumption | change with (a=a); apply rule #;]]]
+qed.
+
+variant eq_cones_to_eq_rel':
+ ∀S,T. ∀f,g: arrows1 REL S T.
+ (∀x:S. or_f ?? (map_arrows2 ?? POW S T f) {(x)} = or_f ?? (map_arrows2 ?? POW S T g) {(x)}) →
+ f = g
+≝ eq_cones_to_eq_rel.
+
+(*
+lemma rOR_full :
+ ∀s,t:rOBP.∀f:arrows2 OBTop (OR (ℱ_2 s)) (OR (ℱ_2 t)).
+ exT22 ? (λg:arrows2 rOBP s t.
+ map_arrows2 ?? OR ?? (ℳ_2 g) = f).
+intros 2 (s t); cases s (s_2 s_1 s_eq); clear s;
+change in match (F2 ??? (mk_Fo ??????)) with s_2;
+cases s_eq; clear s_eq s_2;
+letin s1 ≝ (BP_to_OBP s_1); change in match (BP_to_OBP s_1) with s1;
+cases t (t_2 t_1 t_eq); clear t;
+change in match (F2 ??? (mk_Fo ??????)) with t_2;
+cases t_eq; clear t_eq t_2;
+letin t1 ≝ (BP_to_OBP t_1); change in match (BP_to_OBP t_1) with t1;
+whd in ⊢ (%→?); whd in ⊢ (? (? ? ? ? %) (? ? ? ? %)→?);
+intro; whd in s_1 t_1;
+letin R ≝ (? : (carr2 (arrows2 (category2_of_category1 BP) s_1 t_1)));
+[2:
+ exists;
+ [ constructor 1;
+ [2: simplify; apply R;
+ | simplify; apply (fun12 ?? (map_arrows2 ?? BP_to_OBP s_1 t_1)); apply R;
+ | simplify; apply rule #; ]]
+ simplify;
+|1: constructor 1;
+ [2: apply (pi1exT22 ?? (POW_full (form s_1) (form t_1) f));
+ |1: letin u ≝ (or_f_star ?? (map_arrows2 ?? POW (concr t_1) (form t_1) (⊩ \sub t_1)));
+ letin r ≝ (u ∘ (or_f ?? f));
+ letin xxx ≝ (or_f ?? (map_arrows2 ?? POW (concr s_1) (form s_1) (⊩ \sub s_1)));
+ letin r' ≝ (r ∘ xxx); clearbody r';
+ apply (POW_full' (concr s_1) (concr t_1) r');
+ | simplify in ⊢ (? ? ? (? ? ? ? ? % ?) ?);
+ apply eq_cones_to_eq_rel'; intro;
+ apply
+ (cic:/matita/logic/equality/eq_elim_r''.con ?????
+ (category2_of_category1_respects_comp_r : ?));
+ apply rule (.= (#‡#));
+ apply (.= (respects_comp2 ?? POW (concr s_1) (concr t_1) (form t_1) ? (⊩\sub t_1))‡#);
+ apply sym2;
+ apply (.= (respects_comp2 ?? POW (concr s_1) (form s_1) (form t_1) (⊩\sub s_1) (pi1exT22 ?? (POW_full (form s_1) (form t_1) (Ocont_rel ?? f)))));
+ apply (let H ≝(\snd (POW_full (form s_1) (form t_1) (Ocont_rel ?? f))) in .= #‡H);
+ apply sym2;
+ ]
+
+STOP;
+
+
+(* Todo: rename BTop → OBTop *)
+
+(* scrivo gli statement qua cosi' verra' un conflitto :-)
+
+1. definire il funtore OR
+2. dimostrare che ORel e' faithful
+
+3. Definire la funzione
+ Apply:
+ ∀C1,C2: CAT2. F: arrows3 CAT2 C1 C2 → CAT2
+ ≝
+ constructor 1;
+ [ gli oggetti sono gli oggetti di C1 mappati da F
+ | i morfismi i morfismi di C1 mappati da F
+ | ....
+ ]
+
+ E : objs CATS === Σx.∃y. F y = x
+
+ Quindi (Apply C1 C2 F) (che usando da ora in avanti una coercion
+ scrivero' (F C1) ) e' l'immagine di C1 tramite F ed e'
+ una sottocategoria di C2 (qualcosa da dimostare qui??? vedi sotto
+ al punto 5)
+
+4. Definire rOBP (le OBP rappresentabili) come (BP_to_OBP BP)
+ [Si puo' fare lo stesso per le OA: rOA ≝ Rel_to_OA REL ]
+
+5. Dimostrare che OR (il funtore faithful da OBP a OBTop) e' full
+ quando applicato a rOBP.
+ Nota: puo' darsi che faccia storie ad accettare lo statement.
+ Infatti rOBP e' (BP_to_OBP BP) ed e' "una sottocategoria di OBP"
+ e OR va da OBP a OBTop. Non so se tipa subito o se devi dare
+ una "proiezione" da rOBP a OBP.
+
+6. Definire rOBTop come (OBP_to_OBTop rOBP).
+
+7. Per composizione si ha un funtore full and faithful da BP a rOBTop:
+ basta prendere (OR ∘ BP_to_OBP).
+
+8. Dimostrare (banale: quasi tutti i campi sono per conversione) che
+ esiste un funtore da rOBTop a BTop. Dimostrare che tale funtore e'
+ faithful e full (banale: tutta conversione).
+
+9. Per composizione si ha un funtore full and faithful da BP a BTop.
+
+10. Dimostrare che i seguenti funtori sono anche isomorphism-dense
+ (http://planetmath.org/encyclopedia/DenseFunctor.html):
+
+ BP_to_OBP
+ OBP_to_OBTop quando applicato alle rOBP
+ OBTop_to_BTop quando applicato alle rOBTop
+
+ Concludere per composizione che anche il funtore da BP a BTop e'
+ isomorphism-dense.
+
+====== Da qui in avanti non e' "necessario" nulla:
+
+== altre cose mancanti
+
+11. Dimostrare che le r* e le * orrizzontali
+ sono isomorfe dando il funtore da r* a * e dimostrando che componendo i
+ due funtori ottengo l'identita'
+
+12. La definizione di r* fa schifo: in pratica dici solo come ottieni
+ qualcosa, ma non come lo caratterizzeresti. Ora un teorema carino
+ e' che una a* (e.g. una aOBP) e' sempre una rOBP dove "a" sta per
+ atomic. Dimostrarlo per tutte le r*.
+
+== categorish/future works
+
+13. definire astrattamente la FG-completion e usare quella per
+ ottenere le BP da Rel e le OBP da OA.
+
+14. indebolire le OA, generalizzare le costruzioni, etc. come detto
+ con Giovanni
+
+*)
+
+*)
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/subsets.ma".
+
+record binary_relation (A,B: SET) : Type[1] ≝
+ { satisfy:> binary_morphism1 A B CPROP }.
+
+notation < "hvbox (x \nbsp \natur term 90 r \nbsp y)" with precedence 45 for @{'satisfy $r $x $y}.
+notation > "hvbox (x \natur term 90 r y)" with precedence 45 for @{'satisfy $r $x $y}.
+interpretation "relation applied" 'satisfy r x y = (fun21 ??? (satisfy ?? r) x y).
+
+definition binary_relation_setoid: SET → SET → setoid1.
+ intros (A B);
+ constructor 1;
+ [ apply (binary_relation A B)
+ | constructor 1;
+ [ apply (λA,B.λr,r': binary_relation A B. ∀x,y. r x y ↔ r' x y)
+ | simplify; intros 3; split; intro; assumption
+ | simplify; intros 5; split; intro;
+ [ apply (fi ?? (f ??)) | apply (if ?? (f ??))] assumption
+ | simplify; intros 7; split; intro;
+ [ apply (if ?? (f1 ??)) | apply (fi ?? (f ??)) ]
+ [ apply (if ?? (f ??)) | apply (fi ?? (f1 ??)) ]
+ assumption]]
+qed.
+
+definition binary_relation_of_binary_relation_setoid :
+ ∀A,B.binary_relation_setoid A B → binary_relation A B ≝ λA,B,c.c.
+coercion binary_relation_of_binary_relation_setoid.
+
+definition composition:
+ ∀A,B,C.
+ (binary_relation_setoid A B) × (binary_relation_setoid B C) ⇒_1 (binary_relation_setoid A C).
+ intros;
+ constructor 1;
+ [ intros (R12 R23);
+ constructor 1;
+ constructor 1;
+ [ apply (λs1:A.λs3:C.∃s2:B. s1 ♮R12 s2 ∧ s2 ♮R23 s3);
+ | intros;
+ split; intro; cases e2 (w H3); clear e2; exists; [1,3: apply w ]
+ [ apply (. (e^-1‡#)‡(#‡e1^-1)); assumption
+ | apply (. (e‡#)‡(#‡e1)); assumption]]
+ | intros 8; split; intro H2; simplify in H2 ⊢ %;
+ cases H2 (w H3); clear H2; exists [1,3: apply w] cases H3 (H2 H4); clear H3;
+ [ lapply (if ?? (e x w) H2) | lapply (fi ?? (e x w) H2) ]
+ [ lapply (if ?? (e1 w y) H4)| lapply (fi ?? (e1 w y) H4) ]
+ exists; try assumption;
+ split; assumption]
+qed.
+
+definition REL: category1.
+ constructor 1;
+ [ apply setoid
+ | intros (T T1); apply (binary_relation_setoid T T1)
+ | intros; constructor 1;
+ constructor 1; unfold setoid1_of_setoid; simplify;
+ [ (* changes required to avoid universe inconsistency *)
+ change with (carr o → carr o → CProp); intros; apply (eq ? c c1)
+ | intros; split; intro; change in a a' b b' with (carr o);
+ change in e with (eq ? a a'); change in e1 with (eq ? b b');
+ [ apply (.= (e ^ -1));
+ apply (.= e2);
+ apply e1
+ | apply (.= e);
+ apply (.= e2);
+ apply (e1 ^ -1)]]
+ | apply composition
+ | intros 9;
+ split; intro;
+ cases f (w H); clear f; cases H; clear H;
+ [cases f (w1 H); clear f | cases f1 (w1 H); clear f1]
+ cases H; clear H;
+ exists; try assumption;
+ split; try assumption;
+ exists; try assumption;
+ split; assumption
+ |6,7: intros 5; unfold composition; simplify; split; intro;
+ unfold setoid1_of_setoid in x y; simplify in x y;
+ [1,3: cases e (w H1); clear e; cases H1; clear H1; unfold;
+ [ apply (. (e : eq1 ? x w)‡#); assumption
+ | apply (. #‡(e : eq1 ? w y)^-1); assumption]
+ |2,4: exists; try assumption; split;
+ (* change required to avoid universe inconsistency *)
+ change in x with (carr o1); change in y with (carr o2);
+ first [apply refl | assumption]]]
+qed.
+
+definition setoid_of_REL : objs1 REL → setoid ≝ λx.x.
+coercion setoid_of_REL.
+
+definition binary_relation_setoid_of_arrow1_REL :
+ ∀P,Q. arrows1 REL P Q → binary_relation_setoid P Q ≝ λP,Q,x.x.
+coercion binary_relation_setoid_of_arrow1_REL.
+
+
+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_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 (λs.{x | True});
+ intros; simplify; split; intro; assumption.
+qed.
+
+coercion full_subset.
+
+definition comprehension: ∀b:REL. (b ⇒_1. CPROP) → Ω^b.
+ apply (λb:REL. λP: b ⇒_1 CPROP. {x | P x});
+ intros; simplify;
+ apply (.= †e); apply refl1.
+qed.
+
+interpretation "subset comprehension" 'comprehension s p =
+ (comprehension s (mk_unary_morphism1 ?? p ?)).
+
+definition ext: ∀X,S:REL. (X ⇒_\r1 S) × S ⇒_1 (Ω^X).
+ intros (X S); constructor 1;
+ [ apply (λr:X ⇒_\r1 S.λf:S.{x ∈ X | x ♮r f}); intros; simplify; apply (.= (e‡#)); apply refl1
+ | intros; simplify; split; intros; simplify;
+ [ change with (∀x. x ♮a b → x ♮a' b'); intros;
+ apply (. (#‡e1^-1)); whd in e; apply (if ?? (e ??)); assumption
+ | change with (∀x. x ♮a' b' → x ♮a b); intros;
+ apply (. (#‡e1)); whd in e; apply (fi ?? (e ??));assumption]]
+qed.
+
+definition extS: ∀X,S:REL. ∀r:X ⇒_\r1 S. Ω^S ⇒_1 Ω^X.
+ (* ∃ is not yet a morphism apply (λX,S,r,F.{x ∈ X | ∃a. a ∈ F ∧ x ♮r a});*)
+ intros (X S r); constructor 1;
+ [ intro F; constructor 1; constructor 1;
+ [ apply (λx. x ∈ X ∧ ∃a:S. a ∈ F ∧ x ♮r a);
+ | intros; split; intro; cases f (H1 H2); clear f; split;
+ [ apply (. (e^-1‡#)); assumption
+ |3: apply (. (e‡#)); assumption
+ |2,4: cases H2 (w H3); exists; [1,3: apply w]
+ [ apply (. (#‡(e^-1‡#))); assumption
+ | apply (. (#‡(e‡#))); assumption]]]
+ | intros; split; simplify; intros; cases f; cases e1; split;
+ [1,3: assumption
+ |2,4: exists; [1,3: apply w]
+ [ apply (. (#‡e^-1)‡#); assumption
+ | apply (. (#‡e)‡#); assumption]]]
+qed.
+(*
+lemma equalset_extS_id_X_X: ∀o:REL.∀X.extS ?? (id1 ? o) X = X.
+ intros;
+ unfold extS; simplify;
+ split; simplify;
+ [ intros 2; change with (a ∈ X);
+ cases f; clear f;
+ cases H; clear H;
+ cases x; clear x;
+ change in f2 with (eq1 ? a w);
+ apply (. (f2\sup -1‡#));
+ assumption
+ | intros 2; change in f with (a ∈ X);
+ split;
+ [ whd; exact I
+ | exists; [ apply a ]
+ split;
+ [ assumption
+ | change with (a = a); apply refl]]]
+qed.
+
+lemma extS_com: ∀o1,o2,o3,c1,c2,S. extS o1 o3 (c2 ∘ c1) S = extS o1 o2 c1 (extS o2 o3 c2 S).
+ intros; unfold extS; simplify; split; intros; simplify; intros;
+ [ cases f (H1 H2); cases H2 (w H3); clear f H2; split; [assumption]
+ cases H3 (H4 H5); cases H5 (w1 H6); clear H3 H5; cases H6 (H7 H8); clear H6;
+ exists; [apply w1] split [2: assumption] constructor 1; [assumption]
+ exists; [apply w] split; assumption
+ | cases f (H1 H2); cases H2 (w H3); clear f H2; split; [assumption]
+ cases H3 (H4 H5); cases H4 (w1 H6); clear H3 H4; cases H6 (w2 H7); clear H6;
+ cases H7; clear H7; exists; [apply w2] split; [assumption] exists [apply w] split;
+ assumption]
+qed.
+*)
+
+(* the same as ⋄ for a basic pair *)
+definition image: ∀U,V:REL. (U ⇒_\r1 V) ⇒_2 (Ω^U ⇒_2 Ω^V).
+ intros; constructor 1;
+ [ intro r; constructor 1;
+ [ apply (λS: Ω^U. {y | ∃x:U. x ♮r y ∧ x ∈ S });
+ intros; simplify; split; intro; cases e1; exists [1,3: apply w]
+ [ apply (. (#‡e^-1)‡#); assumption
+ | apply (. (#‡e)‡#); assumption]
+ | intros; split;
+ [ intro y; simplify; intro yA; cases yA; exists; [ apply w ];
+ apply (. #‡(#‡e^-1)); assumption;
+ | intro y; simplify; intro yA; cases yA; exists; [ apply w ];
+ apply (. #‡(#‡e)); assumption;]]
+ | simplify; intros; intro y; simplify; split; simplify; intros (b H); cases H;
+ exists; [1,3: apply w]; cases x; split; try assumption;
+ [ apply (if ?? (e ??)); | apply (fi ?? (e ??)); ] assumption;]
+qed.
+
+(* the same as □ for a basic pair *)
+definition minus_star_image: ∀U,V:REL. (U ⇒_\r1 V) ⇒_2 (Ω^U ⇒_2 Ω^V).
+ intros; constructor 1; intros;
+ [ constructor 1;
+ [ apply (λS: Ω^U. {y | ∀x:U. x ♮c y → x ∈ S});
+ intros; simplify; split; intros; apply f;
+ [ apply (. #‡e); | apply (. #‡e ^ -1)] assumption;
+ | intros; split; intro; simplify; intros;
+ [ apply (. #‡e^-1);| apply (. #‡e); ] apply f; assumption;]
+ | intros; intro; simplify; split; simplify; intros; apply f;
+ [ apply (. (e x a2)); assumption | apply (. (e^-1 x a2)); assumption]]
+qed.
+
+(* the same as Rest for a basic pair *)
+definition star_image: ∀U,V:REL. (U ⇒_\r1 V) ⇒_2 (Ω^V ⇒_2 Ω^U).
+ intros; constructor 1;
+ [ intro r; constructor 1;
+ [ apply (λS: Ω \sup V. {x | ∀y:V. x ♮r y → y ∈ S});
+ intros; simplify; split; intros; apply f;
+ [ apply (. e ‡#);| apply (. e^ -1‡#);] assumption;
+ | intros; split; simplify; intros;
+ [ apply (. #‡e^-1);| apply (. #‡e); ] apply f; assumption;]
+ | intros; intro; simplify; split; simplify; intros; apply f;
+ [ apply (. e a2 y); | apply (. e^-1 a2 y)] assumption;]
+qed.
+
+(* the same as Ext for a basic pair *)
+definition minus_image: ∀U,V:REL. (U ⇒_\r1 V) ⇒_2 (Ω^V ⇒_2 Ω^U).
+ intros; constructor 1;
+ [ intro r; constructor 1;
+ [ apply (λS: Ω^V. {x | ∃y:V. x ♮r y ∧ y ∈ S }).
+ intros; simplify; split; intros; cases e1; cases x; exists; [1,3: apply w]
+ split; try assumption; [ apply (. (e^-1‡#)); | apply (. (e‡#));] assumption;
+ | intros; simplify; split; simplify; intros; cases e1; cases x;
+ exists [1,3: apply w] split; try assumption;
+ [ apply (. (#‡e^-1)); | apply (. (#‡e));] assumption]
+ | intros; intro; simplify; split; simplify; intros; cases e1; exists [1,3: apply w]
+ cases x; split; try assumption;
+ [ apply (. e^-1 a2 w); | apply (. e a2 w)] assumption;]
+qed.
+
+definition foo : ∀o1,o2:REL.carr1 (o1 ⇒_\r1 o2) → carr2 (setoid2_of_setoid1 (o1 ⇒_\r1 o2)) ≝ λo1,o2,x.x.
+
+interpretation "relation f⎻*" 'OR_f_minus_star r = (fun12 ?? (minus_star_image ? ?) (foo ?? r)).
+interpretation "relation f⎻" 'OR_f_minus r = (fun12 ?? (minus_image ? ?) (foo ?? r)).
+interpretation "relation f*" 'OR_f_star r = (fun12 ?? (star_image ? ?) (foo ?? r)).
+
+definition image_coercion: ∀U,V:REL. (U ⇒_\r1 V) → Ω^U ⇒_2 Ω^V.
+intros (U V r Us); apply (image U V r); qed.
+coercion image_coercion.
+
+(* minus_image is the same as ext *)
+
+theorem image_id: ∀o. (id1 REL o : carr2 (Ω^o ⇒_2 Ω^o)) =_1 (id2 SET1 Ω^o).
+ intros; unfold image_coercion; unfold image; simplify;
+ whd in match (?:carr2 ?);
+ intro U; simplify; split; simplify; intros;
+ [ change with (a ∈ U);
+ cases e; cases x; change in e1 with (w =_1 a); apply (. e1^-1‡#); assumption
+ | change in f with (a ∈ U);
+ exists; [apply a] split; [ change with (a = a); apply refl1 | assumption]]
+qed.
+
+theorem minus_star_image_id: ∀o:REL.
+ ((id1 REL o)⎻* : carr2 (Ω^o ⇒_2 Ω^o)) =_1 (id2 SET1 Ω^o).
+ intros; unfold minus_star_image; simplify; intro U; simplify;
+ split; simplify; intros;
+ [ change with (a ∈ U); apply f; change with (a=a); apply refl1
+ | change in f1 with (eq1 ? x a); apply (. f1‡#); apply f]
+qed.
+
+alias symbol "compose" (instance 5) = "category2 composition".
+alias symbol "compose" (instance 4) = "category1 composition".
+theorem image_comp: ∀A,B,C.∀r:B ⇒_\r1 C.∀s:A ⇒_\r1 B.
+ ((r ∘ s) : carr2 (Ω^A ⇒_2 Ω^C)) =_1 r ∘ s.
+ intros; intro U; split; intro x; (unfold image; unfold SET1; simplify);
+ intro H; cases H;
+ cases x1; [cases f|cases f1]; exists; [1,3: apply w1] cases x2; split; try assumption;
+ exists; try assumption; split; assumption;
+qed.
+
+theorem minus_star_image_comp:
+ ∀A,B,C.∀r:B ⇒_\r1 C.∀s:A ⇒_\r1 B.
+ minus_star_image A C (r ∘ s) =_1 minus_star_image B C r ∘ (minus_star_image A B s).
+ intros; unfold minus_star_image; intro X; simplify; split; simplify; intros;
+ [ whd; intros; simplify; whd; intros; apply f; exists; try assumption; split; assumption;
+ | cases f1; cases x1; apply f; assumption]
+qed.
+
+
+(*
+(*CSC: unused! *)
+theorem ext_comp:
+ ∀o1,o2,o3: REL.
+ ∀a: arrows1 ? o1 o2.
+ ∀b: arrows1 ? o2 o3.
+ ∀x. ext ?? (b∘a) x = extS ?? a (ext ?? b x).
+ intros;
+ unfold ext; unfold extS; simplify; split; intro; simplify; intros;
+ cases f; clear f; split; try assumption;
+ [ cases f2; clear f2; cases x1; clear x1; exists; [apply w] split;
+ [1: split] assumption;
+ | cases H; clear H; cases x1; clear x1; exists [apply w]; split;
+ [2: cases f] assumption]
+qed.
+*)
+
+axiom daemon : False.
+
+theorem extS_singleton:
+ ∀o1,o2.∀a.∀x.extS o1 o2 a {(x)} = ext o1 o2 a x.
+ intros; unfold extS; unfold ext; unfold singleton; simplify;
+ split; intros 2; simplify; simplify in f;
+ [ cases f; cases e; cases x1; change in f2 with (x =_1 w); apply (. #‡f2); assumption;
+ | split; try apply I; exists [apply x] split; try assumption; change with (x = x); apply rule #;] qed.
\ No newline at end of file
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/relations.ma".
+include "formal_topology/o-algebra.ma".
+
+definition POW': objs1 SET → OAlgebra.
+ intro A; constructor 1;
+ [ apply (Ω^A);
+ | apply subseteq;
+ | apply overlaps;
+ | apply big_intersects;
+ | apply big_union;
+ | apply ({x | True});
+ simplify; intros; apply (refl1 ? (eq1 CPROP));
+ | apply ({x | False});
+ simplify; intros; apply (refl1 ? (eq1 CPROP));
+ | intros; whd; intros; assumption
+ | intros; whd; split; assumption
+ | intros; apply transitive_subseteq_operator; [2: apply f; | skip | assumption]
+ | intros; cases f; exists [apply w] assumption
+ | intros; split; [ intros 4; apply (f ? f1 i); | intros 3; intro; apply (f i ? f1); ]
+ | intros; split;
+ [ intros 4; apply f; exists; [apply i] assumption;
+ | intros 3; intros; cases f1; apply (f w a x); ]
+ | intros 3; cases f;
+ | intros 3; constructor 1;
+ | intros; cases f; exists; [apply w]
+ [ assumption
+ | whd; intros; cases i; simplify; assumption]
+ | intros; split; intro;
+ [ (** screenshot "screen-pow". *) cases f; cases x1; exists [apply w1] exists [apply w] assumption;
+ | cases e; cases x; exists; [apply w1] [ assumption | exists; [apply w] assumption]]
+ | intros; intros 2; cases (f {(a)} ?);
+ [ exists; [apply a] [assumption | change with (a = a); apply refl1;]
+ | change in x1 with (a = w); change with (mem A a q); apply (. (x1‡#));
+ assumption]]
+qed.
+
+definition powerset_of_POW': ∀A.oa_P (POW' A) → Ω^A ≝ λA,x.x.
+coercion powerset_of_POW'.
+
+definition orelation_of_relation: ∀o1,o2:REL. o1 ⇒_\r1 o2 → (POW' o1) ⇒_\o2 (POW' o2).
+ intros;
+ constructor 1;
+ [ apply rule c;
+ | apply rule (c⎻* );
+ | apply rule (c* );
+ | apply rule (c⎻);
+ | intros; split; intro;
+ [ intros 2; intros 2; apply (f y); exists[apply a] split; assumption;
+ | intros 2; change with (a ∈ q); cases f1; cases x; clear f1 x;
+ apply (f w f3); assumption; ]
+ | unfold foo; intros; split; intro;
+ [ intros 2; intros 2; apply (f x); exists [apply a] split; assumption;
+ | intros 2; change with (a ∈ q); cases f1; cases x; apply (f w f3); assumption;]
+ | intros; split; unfold foo; unfold image_coercion; simplify; intro; cases f; clear f;
+ [ cases x; cases x2; clear x x2; exists; [apply w1]
+ [ assumption | exists; [apply w] split; assumption]
+ | cases x1; cases x2; clear x1 x2; exists; [apply w1]
+ [ exists; [apply w] split; assumption;
+ | assumption; ]]]
+qed.
+
+lemma orelation_of_relation_preserves_equality:
+ ∀o1,o2:REL.∀t,t': o1 ⇒_\r1 o2.
+ t = t' → orelation_of_relation ?? t =_2 orelation_of_relation ?? t'.
+ intros; split; unfold orelation_of_relation; unfold foo; simplify;
+ change in e with (t =_2 t'); unfold image_coercion; apply (†e);
+qed.
+
+lemma minus_image_id : ∀o:REL.((id1 REL o))⎻ =_1 (id2 SET1 Ω^o).
+unfold foo; intro o; intro; unfold minus_image; simplify; split; simplify; intros;
+[ cases e; cases x; change with (a1 ∈ a); change in f with (a1 =_1 w);
+ apply (. f‡#); assumption;
+| change in f with (a1 ∈ a); exists [ apply a1] split; try assumption;
+ change with (a1 =_1 a1); apply refl1;]
+qed.
+
+lemma star_image_id : ∀o:REL. ((id1 REL o))* =_1 (id2 SET1 Ω^o).
+unfold foo; intro o; intro; unfold star_image; simplify; split; simplify; intros;
+[ change with (a1 ∈ a); apply f; change with (a1 =_1 a1); apply rule refl1;
+| change in f1 with (a1 =_1 y); apply (. f1^-1‡#); apply f;]
+qed.
+
+lemma orelation_of_relation_preserves_identity:
+ ∀o1:REL. orelation_of_relation ?? (id1 ? o1) =_2 id2 OA (POW' o1).
+ intros; split;
+ (unfold orelation_of_relation; unfold OA; unfold foo; simplify);
+ [ apply (minus_star_image_id o1);
+ | apply (minus_image_id o1);
+ | apply (image_id o1);
+ | apply (star_image_id o1) ]
+qed.
+
+(*
+ split; whd; intro;
+ [ change with ((∀x. x ♮(id1 REL o1) a1→x∈a) → a1 ∈ a); intros;
+ apply (f a1); change with (a1 = a1); apply refl1;
+ | change with (a1 ∈ a → ∀x. x ♮(id1 REL o1) a1→x∈a); intros;
+ change in f1 with (x = a1); apply (. f1‡#); apply f;
+ | alias symbol "and" = "and_morphism".
+ change with ((∃y:o1.a1 ♮(id1 REL o1) y ∧ y∈a) → a1 ∈ a);
+ intro; cases e; clear e; cases x; clear x; change in f with (a1=w);
+ apply (. f‡#); apply f1;
+ | change with (a1 ∈ a → ∃y:o1.a1 ♮(id1 REL o1) y ∧ y∈a);
+ intro; exists; [apply a1]; split; [ change with (a1=a1); apply refl1; | apply f]
+ | change with ((∃x:o1.x ♮(id1 REL o1) a1∧x∈a) → a1 ∈ a);
+ intro; cases e; clear e; cases x; clear x; change in f with (w=a1);
+ apply (. f^-1‡#); apply f1;
+ | change with (a1 ∈ a → ∃x:o1.x ♮(id1 REL o1) a1∧x∈a);
+ intro; exists; [apply a1]; split; [ change with (a1=a1); apply refl1; | apply f]
+ | change with ((∀y.a1 ♮(id1 REL o1) y→y∈a) → a1 ∈ a); intros;
+ apply (f a1); change with (a1 = a1); apply refl1;
+ | change with (a1 ∈ a → ∀y.a1 ♮(id1 REL o1) y→y∈a); intros;
+ change in f1 with (a1 = y); apply (. f1^-1‡#); apply f;]
+qed.
+*)
+
+(* CSC: ???? forse un uncertain mancato *)
+alias symbol "eq" = "setoid2 eq".
+alias symbol "compose" = "category1 composition".
+lemma orelation_of_relation_preserves_composition:
+ ∀o1,o2,o3:REL.∀F: o1 ⇒_\r1 o2.∀G: o2 ⇒_\r1 o3.
+ orelation_of_relation ?? (G ∘ F) =
+ comp2 OA ??? (orelation_of_relation ?? F) (orelation_of_relation ?? G).
+ intros; split; intro; split; whd; intro; whd in ⊢ (% → %); intros;
+ [ whd; intros; apply f; exists; [ apply x] split; assumption;
+ | cases f1; clear f1; cases x1; clear x1; apply (f w); assumption;
+ | cases e; cases x; cases f; cases x1; clear e x f x1; exists; [ apply w1 ]
+ split; [ assumption | exists; [apply w] split; assumption ]
+ | cases e; cases x; cases f1; cases x1; clear e x f1 x1; exists; [apply w1 ]
+ split; [ exists; [apply w] split; assumption | assumption ]
+ | unfold arrows1_of_ORelation_setoid;
+ cases e; cases x; cases f; cases x1; clear e x f x1; exists; [ apply w1 ]
+ split; [ assumption | exists; [apply w] split; assumption ]
+ | unfold arrows1_of_ORelation_setoid in e;
+ cases e; cases x; cases f1; cases x1; clear e x f1 x1; exists; [apply w1 ]
+ split; [ exists; [apply w] split; assumption | assumption ]
+ | whd; intros; apply f; exists; [ apply y] split; assumption;
+ | cases f1; clear f1; cases x; clear x; apply (f w); assumption;]
+qed.
+
+definition POW: carr3 (arrows3 CAT2 (category2_of_category1 REL) OA).
+ constructor 1;
+ [ apply POW';
+ | intros; constructor 1;
+ [ apply (orelation_of_relation S T);
+ | intros; apply (orelation_of_relation_preserves_equality S T a a' e); ]
+ | apply orelation_of_relation_preserves_identity;
+ | apply orelation_of_relation_preserves_composition; ]
+qed.
+
+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; simplify; unfold image_coercion in e3; cases (e3 {(x)});
+ split; intro; [ lapply (s y); | lapply (s1 y); ]
+ [2,4: exists; [1,3:apply x] split; [1,3: assumption |*: change with (x=x); apply rule #]
+ |*: 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.
+*)
+
+include "formal_topology/notation.ma".
+
+theorem POW_full: full2 ?? POW.
+ intros 3 (S T); exists;
+ [ constructor 1; constructor 1;
+ [ apply (λx:carr S.λy:carr T. y ∈ f {(x)});
+ | apply hide; intros; unfold FunClass_1_OF_carr2; lapply (.= e1‡#);
+ [4: apply mem; |6: apply Hletin;|1,2,3,5: skip]
+ lapply (#‡prop11 ?? f ?? (†e)); [6: apply Hletin; |*:skip ]]
+ | (split; intro; split; simplify);
+ [ change with (∀a1.(∀x. a1 ∈ (f {(x):S}) → x ∈ a) → a1 ∈ f⎻* a);
+ | change with (∀a1.a1 ∈ f⎻* a → (∀x.a1 ∈ f {(x):S} → x ∈ a));
+ | alias symbol "and" (instance 4) = "and_morphism".
+change with (∀a1.(∃y:carr T. y ∈ f {(a1):S} ∧ y ∈ a) → a1 ∈ f⎻ a);
+ | alias symbol "and" (instance 2) = "and_morphism".
+change with (∀a1.a1 ∈ f⎻ a → (∃y:carr T.y ∈ f {(a1):S} ∧ y ∈ a));
+ | alias symbol "and" (instance 3) = "and_morphism".
+change with (∀a1.(∃x:carr S. a1 ∈ f {(x):S} ∧ x ∈ a) → a1 ∈ f a);
+ | change with (∀a1.a1 ∈. f a → (∃x:carr S. a1 ∈ f {(x):S} ∧ x ∈ a));
+ | change with (∀a1.(∀y. y ∈ f {(a1):S} → y ∈ a) → a1 ∈ f* a);
+ | change with (∀a1.a1 ∈ f* a → (∀y. y ∈ f {(a1):S} → y ∈ a)); ]
+ [ intros; apply ((. (or_prop2 ?? f (singleton ? a1) a)^-1) ? a1);
+ [ intros 2; apply (f1 a2); change in f2 with (a2 ∈ f⎻ (singleton ? a1));
+ lapply (. (or_prop3 ?? f (singleton ? a2) (singleton ? a1)));
+ [ cases Hletin; change in x1 with (eq1 ? a1 w);
+ apply (. x1‡#); assumption;
+ | exists; [apply a2] [change with (a2=a2); apply rule #; | assumption]]
+ | change with (a1 = a1); apply rule #; ]
+ | intros; apply ((. (or_prop2 ?? f (singleton ? a1) a)) ? x);
+ [ intros 2; change in f3 with (eq1 ? a1 a2); change with (a2 ∈ f⎻* a); apply (. f3^-1‡#);
+ assumption;
+ | lapply (. (or_prop3 ?? f (singleton ? x) (singleton ? a1))^-1);
+ [ cases Hletin; change in x1 with (eq1 ? x w);
+ change with (x ∈ f⎻ (singleton ? a1)); apply (. x1‡#); assumption;
+ | exists; [apply a1] [assumption | change with (a1=a1); apply rule #; ]]]
+ | intros; cases e; cases x; clear e x;
+ lapply (. (or_prop3 ?? f (singleton ? a1) a)^-1);
+ [ cases Hletin; change in x with (eq1 ? a1 w1); apply (. x‡#); assumption;
+ | exists; [apply w] assumption ]
+ | intros; lapply (. (or_prop3 ?? f (singleton ? a1) a));
+ [ cases Hletin; exists; [apply w] split; assumption;
+ | exists; [apply a1] [change with (a1=a1); apply rule #; | assumption ]]
+ | intros; cases e; cases x; clear e x;
+ apply (f_image_monotone ?? f (singleton ? w) a ? a1);
+ [ intros 2; change in f3 with (eq1 ? w a2); change with (a2 ∈ a);
+ apply (. f3^-1‡#); assumption;
+ | assumption; ]
+ | intros; lapply (. (or_prop3 ?? f a (singleton ? a1))^-1);
+ [ cases Hletin; exists; [apply w] split;
+ [ lapply (. (or_prop3 ?? f (singleton ? w) (singleton ? a1)));
+ [ cases Hletin1; change in x3 with (eq1 ? a1 w1); apply (. x3‡#); assumption;
+ | exists; [apply w] [change with (w=w); apply rule #; | assumption ]]
+ | assumption ]
+ | exists; [apply a1] [ assumption; | change with (a1=a1); apply rule #;]]
+ | intros; apply ((. (or_prop1 ?? f (singleton ? a1) a)^-1) ? a1);
+ [ apply f1; | change with (a1=a1); apply rule #; ]
+ | intros; apply ((. (or_prop1 ?? f (singleton ? a1) a)) ? y);
+ [ intros 2; change in f3 with (eq1 ? a1 a2); change with (a2 ∈ f* a);
+ apply (. f3^-1‡#); assumption;
+ | assumption ]]]
+qed.
\ No newline at end of file
--- /dev/null
+#!/bin/sh
+for MA in `find -name "*.ma"`; do
+ echo ${MA}; sed "s!$1!$2!g" ${MA} > ${MA}.new
+ if diff ${MA} ${MA}.new > /dev/null;
+ then rm -f ${MA}.new;
+ else mv -f ${MA} ${MA}.old; mv -f ${MA}.new ${MA};
+ fi
+done
+
+unset MA
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/relations.ma".
+
+definition is_saturation: ∀C:REL. Ω^C ⇒_1 Ω^C → CProp[1] ≝
+ λC:REL.λA:Ω^C ⇒_1 Ω^C. ∀U,V. (U ⊆ A V) =_1 (A U ⊆ A V).
+
+definition is_reduction: ∀C:REL. Ω^C ⇒_1 Ω^C → CProp[1] ≝
+ λC:REL.λJ: Ω^C ⇒_1 Ω^C. ∀U,V. (J U ⊆ V) =_1 (J U ⊆ J V).
+
+theorem saturation_expansive: ∀C,A. is_saturation C A → ∀U. U ⊆ A U.
+ intros; apply (fi ?? (i ??)); apply subseteq_refl.
+qed.
+
+theorem saturation_monotone:
+ ∀C,A. is_saturation C A →
+ ∀U,V. U ⊆ V → A U ⊆ A V.
+ intros; apply (if ?? (i ??)); apply subseteq_trans; [apply V|3: apply saturation_expansive ]
+ assumption.
+qed.
+
+theorem saturation_idempotent: ∀C,A. is_saturation C A → ∀U. A (A U) = A U.
+ intros; split;
+ [ apply (if ?? (i ??)); apply subseteq_refl
+ | apply saturation_expansive; assumption]
+qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/saturations.ma".
+include "formal_topology/o-saturations.ma".
+include "formal_topology/relations_to_o-algebra.ma".
+
+(* These are only conversions :-) *)
+
+definition o_operator_of_operator: ∀C:REL. (Ω^C ⇒_1 Ω^C) → ((POW C) ⇒_1 (POW C)) ≝ λC,t.t.
+
+definition is_o_saturation_of_is_saturation:
+ ∀C:REL.∀R: Ω^C ⇒_1 Ω^C. is_saturation ? R → is_o_saturation ? (o_operator_of_operator ? R).
+intros (C R i); apply i; qed.
+
+definition is_o_reduction_of_is_reduction:
+ ∀C:REL.∀R: Ω^C ⇒_1 Ω^C.is_reduction ? R → is_o_reduction ? (o_operator_of_operator ? R).
+intros (C R i); apply i; qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "formal_topology/categories.ma".
+
+record powerset_carrier (A: objs1 SET) : Type[1] ≝ { mem_operator: A ⇒_1 CPROP }.
+interpretation "powerset low" 'powerset A = (powerset_carrier A).
+notation "hvbox(a break ∈. b)" non associative with precedence 45 for @{ 'mem_low $a $b }.
+interpretation "memlow" 'mem_low a S = (fun11 ?? (mem_operator ? S) a).
+
+definition subseteq_operator: ∀A: objs1 SET. Ω^A → Ω^A → CProp[0] ≝
+ λA:objs1 SET.λU,V.∀a:A. a ∈. U → a ∈. V.
+
+theorem transitive_subseteq_operator: ∀A. transitive2 ? (subseteq_operator A).
+ intros 6; intros 2;
+ apply s1; apply s;
+ assumption.
+qed.
+
+definition powerset_setoid1: SET → SET1.
+ intros (T);
+ constructor 1;
+ [ apply (powerset_carrier T)
+ | constructor 1;
+ [ apply (λU,V. subseteq_operator ? U V ∧ subseteq_operator ? V U)
+ | simplify; intros; split; intros 2; assumption
+ | simplify; intros (x y H); cases H; split; assumption
+ | simplify; intros (x y z H H1); cases H; cases H1; split;
+ apply transitive_subseteq_operator; [1,4: apply y ]
+ assumption ]]
+qed.
+
+interpretation "powerset" 'powerset A = (powerset_setoid1 A).
+
+interpretation "subset construction" 'subset \eta.x =
+ (mk_powerset_carrier ? (mk_unary_morphism1 ? CPROP x ?)).
+
+definition mem: ∀A. A × Ω^A ⇒_1 CPROP.
+ intros;
+ constructor 1;
+ [ apply (λx,S. mem_operator ? S x)
+ | intros 5;
+ cases b; clear b; cases b'; clear b'; simplify; intros;
+ apply (trans1 ????? (prop11 ?? u ?? e));
+ cases e1; whd in s s1;
+ split; intro;
+ [ apply s; assumption
+ | apply s1; assumption]]
+qed.
+
+interpretation "mem" 'mem a S = (fun21 ??? (mem ?) a S).
+
+definition subseteq: ∀A. Ω^A × Ω^A ⇒_1 CPROP.
+ intros;
+ constructor 1;
+ [ apply (λU,V. subseteq_operator ? U V)
+ | intros;
+ cases e; cases e1;
+ split; intros 1;
+ [ apply (transitive_subseteq_operator ????? s2);
+ apply (transitive_subseteq_operator ???? s1 s4)
+ | apply (transitive_subseteq_operator ????? s3);
+ apply (transitive_subseteq_operator ???? s s4) ]]
+qed.
+
+interpretation "subseteq" 'subseteq U V = (fun21 ??? (subseteq ?) U V).
+
+theorem subseteq_refl: ∀A.∀S:Ω^A.S ⊆ S.
+ intros 4; assumption.
+qed.
+
+theorem subseteq_trans: ∀A.∀S1,S2,S3: Ω^A. S1 ⊆ S2 → S2 ⊆ S3 → S1 ⊆ S3.
+ intros; apply transitive_subseteq_operator; [apply S2] assumption.
+qed.
+
+definition overlaps: ∀A. Ω^A × Ω^A ⇒_1 CPROP.
+ intros;
+ constructor 1;
+ [ apply (λA:objs1 SET.λU,V:Ω^A.(exT2 ? (λx:A.x ∈ U) (λx:A.x ∈ V) : CProp[0]))
+ | intros;
+ constructor 1; intro; cases e2; exists; [1,4: apply w]
+ [ apply (. #‡e^-1); assumption
+ | apply (. #‡e1^-1); assumption
+ | apply (. #‡e); assumption;
+ | apply (. #‡e1); assumption]]
+qed.
+
+interpretation "overlaps" 'overlaps U V = (fun21 ??? (overlaps ?) U V).
+
+definition intersects: ∀A. Ω^A × Ω^A ⇒_1 Ω^A.
+ intros;
+ constructor 1;
+ [ apply rule (λU,V. {x | x ∈ U ∧ x ∈ V });
+ intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1;
+ | intros;
+ split; intros 2; simplify in f ⊢ %;
+ [ apply (. (#‡e^-1)‡(#‡e1^-1)); assumption
+ | apply (. (#‡e)‡(#‡e1)); assumption]]
+qed.
+
+interpretation "intersects" 'intersects U V = (fun21 ??? (intersects ?) U V).
+
+definition union: ∀A. Ω^A × Ω^A ⇒_1 Ω^A.
+ intros;
+ constructor 1;
+ [ apply (λU,V. {x | x ∈ U ∨ x ∈ V });
+ intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1
+ | intros;
+ split; intros 2; simplify in f ⊢ %;
+ [ apply (. (#‡e^-1)‡(#‡e1^-1)); assumption
+ | apply (. (#‡e)‡(#‡e1)); assumption]]
+qed.
+
+interpretation "union" 'union U V = (fun21 ??? (union ?) U V).
+
+(* qua non riesco a mettere set *)
+definition singleton: ∀A:setoid. A ⇒_1 Ω^A.
+ intros; constructor 1;
+ [ apply (λa:A.{b | a =_0 b}); unfold setoid1_of_setoid; simplify;
+ intros; simplify;
+ split; intro;
+ apply (.= e1);
+ [ apply e | apply (e^-1) ]
+ | unfold setoid1_of_setoid; simplify;
+ intros; split; intros 2; simplify in f ⊢ %; apply trans;
+ [ apply a |4: apply a'] try assumption; apply sym; assumption]
+qed.
+
+interpretation "singleton" 'singl a = (fun11 ?? (singleton ?) a).
+notation > "{ term 19 a : S }" with precedence 90 for @{fun11 ?? (singleton $S) $a}.
+
+definition big_intersects: ∀A:SET.∀I:SET.(I ⇒_2 Ω^A) ⇒_2 Ω^A.
+ intros; constructor 1;
+ [ intro; whd; whd in I;
+ apply ({x | ∀i:I. x ∈ c i});
+ simplify; intros; split; intros; [ apply (. (e^-1‡#)); | apply (. e‡#); ]
+ apply f;
+ | intros; split; intros 2; simplify in f ⊢ %; intro;
+ [ apply (. (#‡(e i)^-1)); apply f;
+ | apply (. (#‡e i)); apply f]]
+qed.
+
+definition big_union: ∀A:SET.∀I:SET.(I ⇒_2 Ω^A) ⇒_2 Ω^A.
+ intros; constructor 1;
+ [ intro; whd; whd in A; whd in I;
+ apply ({x | ∃i:I. x ∈ c i });
+ simplify; intros; split; intros; cases e1; clear e1; exists; [1,3:apply w]
+ [ apply (. (e^-1‡#)); | apply (. (e‡#)); ]
+ apply x;
+ | intros; split; intros 2; simplify in f ⊢ %; cases f; clear f; exists; [1,3:apply w]
+ [ apply (. (#‡(e w)^-1)); apply x;
+ | apply (. (#‡e w)); apply x]]
+qed.
+
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋃) \below (\emsp) term 90 p)"
+non associative with precedence 50 for @{ 'bigcup $p }.
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋃) \below (ident i ∈ I) break term 90 p)"
+non associative with precedence 50 for @{ 'bigcup_mk (λ${ident i}:$I.$p) }.
+notation > "hovbox(⋃ f)" non associative with precedence 60 for @{ 'bigcup $f }.
+
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋂) \below (\emsp) term 90 p)"
+non associative with precedence 50 for @{ 'bigcap $p }.
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋂) \below (ident i ∈ I) break term 90 p)"
+non associative with precedence 50 for @{ 'bigcap_mk (λ${ident i}:$I.$p) }.
+notation > "hovbox(⋂ f)" non associative with precedence 60 for @{ 'bigcap $f }.
+
+interpretation "bigcup" 'bigcup f = (fun12 ?? (big_union ??) f).
+interpretation "bigcap" 'bigcap f = (fun12 ?? (big_intersects ??) f).
+interpretation "bigcup mk" 'bigcup_mk f = (fun12 ?? (big_union ??) (mk_unary_morphism2 ?? f ?)).
+interpretation "bigcap mk" 'bigcap_mk f = (fun12 ?? (big_intersects ??) (mk_unary_morphism2 ?? f ?)).
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/categories.ma".
-include "formal_topology/notation.ma".
-
-record Fo (C1,C2:CAT2) (F:arrows3 CAT2 C1 C2) : Type2 ≝ {
- F2: C2;
- F1: C1;
- FP: map_objs2 ?? F F1 =_\ID F2
-}.
-
-notation "ℱ\sub 1 x" non associative with precedence 60 for @{'F1 $x}.
-notation > "ℱ_1" non associative with precedence 90 for @{F1 ???}.
-interpretation "F1" 'F1 x = (F1 ??? x).
-
-notation "ℱ\sub 2 x" non associative with precedence 60 for @{'F2 $x}.
-notation > "ℱ_2" non associative with precedence 90 for @{F2 ???}.
-interpretation "F2" 'F2 x = (F2 ??? x).
-
-lemma REW : ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.∀X,Y:Fo ?? F.
- arrows2 C2 (F (ℱ_1 X)) (F (ℱ_1 Y)) →
- arrows2 C2 (ℱ_2 X) (ℱ_2 Y).
-intros 5; cases X; cases Y; clear X Y;
-cases H; cases H1; intros; assumption;
-qed.
-
-record Fm_c (C1,C2:CAT2) (F:arrows3 CAT2 C1 C2) (X,Y:Fo ?? F) : Type2 ≝ {
- Fm2: arrows2 C2 (F2 ??? X) (F2 ??? Y);
- Fm1: arrows2 C1 (F1 ??? X) (F1 ??? Y);
- FmP: REW ?? F X Y (map_arrows2 ?? F ?? Fm1) = Fm2
-}.
-
-notation "ℳ\sub 1 x" non associative with precedence 60 for @{'Fm1 $x}.
-notation > "ℳ_1" non associative with precedence 90 for @{Fm1 ?????}.
-interpretation "Fm1" 'Fm1 x = (Fm1 ????? x).
-
-notation "ℳ\sub 2 x" non associative with precedence 60 for @{'Fm2 $x}.
-notation > "ℳ_2" non associative with precedence 90 for @{Fm2 ?????}.
-interpretation "Fm2" 'Fm2 x = (Fm2 ????? x).
-
-definition Fm :
- ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.
- Fo ?? F → Fo ?? F → setoid2.
-intros (C1 C2 F X Y); constructor 1; [apply (Fm_c C1 C2 F X Y)]
-constructor 1; [apply (λf,g.Fm2 ????? f =_2 Fm2 ????? g);]
-[ intro; apply refl2;
-| intros 3; apply sym2; assumption;
-| intros 5; apply (trans2 ?? ??? x1 x2);]
-qed.
-
-definition F_id :
- ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.∀o.Fm ?? F o o.
-intros; constructor 1;
- [ apply (id2 C2 (F2 ??? o));
- | apply (id2 C1 (F1 ??? o));
- | cases o; cases H; simplify; apply (respects_id2 ?? F);]
-qed.
-
-definition F_comp :
- ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.∀o1,o2,o3.
- (Fm ?? F o1 o2) × (Fm ?? F o2 o3) ⇒_2 (Fm ?? F o1 o3).
-intros; constructor 1;
-[ intros (f g); constructor 1;
- [ apply (comp2 C2 ??? (ℳ_2 f) (ℳ_2 g));
- | apply (comp2 C1 ??? (ℳ_1 f) (ℳ_1 g));
- | apply hide; cases o1 in f; cases o2 in g; cases o3; clear o1 o2 o3;
- cases H; cases H1; cases H2; intros 2; cases c; cases c1; clear c c1;
- simplify; apply (.= (respects_comp2:?)); apply (e1‡e);]
-| intros 6; change with ((ℳ_2 b ∘ ℳ_2 a) = (ℳ_2 b' ∘ ℳ_2 a'));
- change in e1 with (ℳ_2 b = ℳ_2 b');
- change in e with (ℳ_2 a = ℳ_2 a');
- apply (e‡e1);]
-qed.
-
-
-definition Apply : ∀C1,C2: CAT2.arrows3 CAT2 C1 C2 → CAT2.
-intros (C1 C2 F);
-constructor 1;
-[ apply (Fo ?? F);
-| apply (Fm ?? F);
-| apply F_id;
-| apply F_comp;
-| intros; apply (comp_assoc2 C2 ???? (ℳ_2 a12) (ℳ_2 a23) (ℳ_2 a34));
-| intros; apply (id_neutral_right2 C2 ?? (ℳ_2 a));
-| intros; apply (id_neutral_left2 C2 ?? (ℳ_2 a));]
-qed.
-
-definition faithful ≝
- λC1,C2.λF:arrows3 CAT2 C1 C2.∀S,T.∀f,g:arrows2 C1 S T.
- map_arrows2 ?? F ?? f = map_arrows2 ?? F ?? g → f=g.
-
-definition Ylppa : ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.
- faithful ?? F → let rC2 ≝ Apply ?? F in arrows3 CAT2 rC2 C1.
-intros; constructor 1;
-[ intro; apply (ℱ_1 o);
-| intros; constructor 1;
- [ intros; apply (ℳ_1 c);
- | apply hide; intros; apply f; change in e with (ℳ_2 a = ℳ_2 a');
- lapply (FmP ????? a) as H1; lapply (FmP ????? a') as H2;
- cut (REW ????? (map_arrows2 ?? F ?? (ℳ_1 a)) =
- REW ????? (map_arrows2 ?? F ?? (ℳ_1 a')));[2:
- apply (.= H1); apply (.= e); apply (H2^-1);]
- clear H1 H2 e; cases S in a a' Hcut; cases T;
- cases H; cases H1; simplify; intros; assumption;]
-| intro; apply rule #;
-| intros; simplify; apply rule #;]
-qed.
-
-
-
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/relations.ma".
-include "formal_topology/notation.ma".
-
-record basic_pair: Type1 ≝ {
- concr: REL; form: REL; rel: concr ⇒_\r1 form
-}.
-
-interpretation "basic pair relation" 'Vdash2 x y c = (fun21 ??? (rel c) x y).
-interpretation "basic pair relation (non applied)" 'Vdash c = (rel c).
-
-record relation_pair (BP1,BP2: basic_pair): Type1 ≝ {
- concr_rel: (concr BP1) ⇒_\r1 (concr BP2); form_rel: (form BP1) ⇒_\r1 (form BP2);
- commute: comp1 REL ??? concr_rel (rel ?) =_1 form_rel ∘ ⊩
- }.
-
-interpretation "concrete relation" 'concr_rel r = (concr_rel ?? r).
-interpretation "formal relation" 'form_rel r = (form_rel ?? r).
-
-definition relation_pair_equality: ∀o1,o2. equivalence_relation1 (relation_pair o1 o2).
- intros; constructor 1; [ apply (λr,r'. ⊩ ∘ r \sub\c = ⊩ ∘ r' \sub\c);
- | simplify; intros; apply refl1;
- | simplify; intros 2; apply sym1;
- | simplify; intros 3; apply trans1; ]
-qed.
-
-definition relation_pair_setoid: basic_pair → basic_pair → setoid1.
- intros;
- constructor 1;
- [ apply (relation_pair b b1)
- | apply relation_pair_equality
- ]
-qed.
-
-definition relation_pair_of_relation_pair_setoid :
- ∀P,Q. relation_pair_setoid P Q → relation_pair P Q ≝ λP,Q,x.x.
-coercion relation_pair_of_relation_pair_setoid.
-
-alias symbol "compose" (instance 1) = "category1 composition".
-lemma eq_to_eq':
- ∀o1,o2.∀r,r':relation_pair_setoid o1 o2. r =_1 r' → r \sub\f ∘ ⊩ =_1 r'\sub\f ∘ ⊩.
- intros 5 (o1 o2 r r' H);
- apply (.= (commute ?? r)^-1);
- change in H with (⊩ ∘ r \sub\c = ⊩ ∘ r' \sub\c);
- apply rule (.= H);
- apply (commute ?? r').
-qed.
-
-definition id_relation_pair: ∀o:basic_pair. relation_pair o o.
- intro;
- constructor 1;
- [1,2: apply id1;
- | lapply (id_neutral_right1 ? (concr o) ? (⊩)) as H;
- lapply (id_neutral_left1 ?? (form o) (⊩)) as H1;
- apply (.= H);
- apply (H1^-1);]
-qed.
-
-lemma relation_pair_composition:
- ∀o1,o2,o3: basic_pair.
- relation_pair_setoid o1 o2 → relation_pair_setoid o2 o3 → relation_pair_setoid o1 o3.
-intros 3 (o1 o2 o3);
- intros (r r1);
- constructor 1;
- [ apply (r1 \sub\c ∘ r \sub\c)
- | apply (r1 \sub\f ∘ r \sub\f)
- | lapply (commute ?? r) as H;
- lapply (commute ?? r1) as H1;
- alias symbol "trans" = "trans1".
- alias symbol "assoc" = "category1 assoc".
- apply (.= ASSOC);
- apply (.= #‡H1);
- alias symbol "invert" = "setoid1 symmetry".
- apply (.= ASSOC ^ -1);
- apply (.= H‡#);
- apply ASSOC]
-qed.
-
-lemma relation_pair_composition_is_morphism:
- ∀o1,o2,o3: basic_pair.
- ∀a,a':relation_pair_setoid o1 o2.
- ∀b,b':relation_pair_setoid o2 o3.
- a=a' → b=b' →
- relation_pair_composition o1 o2 o3 a b
- = relation_pair_composition o1 o2 o3 a' b'.
-intros 3 (o1 o2 o3);
- intros;
- change with (⊩ ∘ (b\sub\c ∘ a\sub\c) = ⊩ ∘ (b'\sub\c ∘ a'\sub\c));
- change in e with (⊩ ∘ a \sub\c = ⊩ ∘ a' \sub\c);
- change in e1 with (⊩ ∘ b \sub\c = ⊩ ∘ b' \sub\c);
- apply (.= ASSOC);
- apply (.= #‡e1);
- apply (.= #‡(commute ?? b'));
- apply (.= ASSOC ^ -1);
- apply (.= e‡#);
- apply (.= ASSOC);
- apply (.= #‡(commute ?? b')^-1);
- apply (ASSOC ^ -1);
-qed.
-
-definition relation_pair_composition_morphism:
- ∀o1,o2,o3. binary_morphism1 (relation_pair_setoid o1 o2) (relation_pair_setoid o2 o3) (relation_pair_setoid o1 o3).
- intros;
- constructor 1;
- [ apply relation_pair_composition;
- | apply relation_pair_composition_is_morphism;]
-qed.
-
-lemma relation_pair_composition_morphism_assoc:
-Πo1:basic_pair
-.Πo2:basic_pair
- .Πo3:basic_pair
- .Πo4:basic_pair
- .Πa12:relation_pair_setoid o1 o2
- .Πa23:relation_pair_setoid o2 o3
- .Πa34:relation_pair_setoid o3 o4
- .relation_pair_composition_morphism o1 o3 o4
- (relation_pair_composition_morphism o1 o2 o3 a12 a23) a34
- =relation_pair_composition_morphism o1 o2 o4 a12
- (relation_pair_composition_morphism o2 o3 o4 a23 a34).
- intros;
- change with (⊩ ∘ (a34\sub\c ∘ (a23\sub\c ∘ a12\sub\c)) =
- ⊩ ∘ ((a34\sub\c ∘ a23\sub\c) ∘ a12\sub\c));
- alias symbol "refl" = "refl1".
- alias symbol "prop2" = "prop21".
- apply (ASSOC‡#);
-qed.
-
-lemma relation_pair_composition_morphism_respects_id:
- ∀o1,o2:basic_pair.∀a:relation_pair_setoid o1 o2.
- relation_pair_composition_morphism o1 o1 o2 (id_relation_pair o1) a=a.
- intros;
- change with (⊩ ∘ (a\sub\c ∘ (id_relation_pair o1)\sub\c) = ⊩ ∘ a\sub\c);
- apply ((id_neutral_right1 ????)‡#);
-qed.
-
-lemma relation_pair_composition_morphism_respects_id_r:
- ∀o1,o2:basic_pair.∀a:relation_pair_setoid o1 o2.
- relation_pair_composition_morphism o1 o2 o2 a (id_relation_pair o2)=a.
- intros;
- change with (⊩ ∘ ((id_relation_pair o2)\sub\c ∘ a\sub\c) = ⊩ ∘ a\sub\c);
- apply ((id_neutral_left1 ????)‡#);
-qed.
-
-definition BP: category1.
- constructor 1;
- [ apply basic_pair
- | apply relation_pair_setoid
- | apply id_relation_pair
- | apply relation_pair_composition_morphism
- | apply relation_pair_composition_morphism_assoc;
- | apply relation_pair_composition_morphism_respects_id;
- | apply relation_pair_composition_morphism_respects_id_r;]
-qed.
-
-definition basic_pair_of_BP : objs1 BP → basic_pair ≝ λx.x.
-coercion basic_pair_of_BP.
-
-definition relation_pair_setoid_of_arrows1_BP :
- ∀P,Q. arrows1 BP P Q → relation_pair_setoid P Q ≝ λP,Q,x.x.
-coercion relation_pair_setoid_of_arrows1_BP.
-
-(*
-definition BPext: ∀o: BP. (form o) ⇒_1 Ω^(concr o).
- intros; constructor 1;
- [ apply (ext ? ? (rel o));
- | intros;
- apply (.= #‡e);
- apply refl1]
-qed.
-
-definition BPextS: ∀o: BP. Ω^(form o) ⇒_1 Ω^(concr o).
- intros; constructor 1;
- [ apply (minus_image ?? (rel o));
- | intros; apply (#‡e); ]
-qed.
-
-definition fintersects: ∀o: BP. (form o) × (form o) ⇒_1 Ω^(form o).
- intros (o); constructor 1;
- [ apply (λa,b: form o.{c | BPext o c ⊆ BPext o a ∩ BPext o b });
- intros; simplify; apply (.= (†e)‡#); apply refl1
- | intros; split; simplify; intros;
- [ apply (. #‡((†e^-1)‡(†e1^-1))); assumption
- | apply (. #‡((†e)‡(†e1))); assumption]]
-qed.
-
-interpretation "fintersects" 'fintersects U V = (fun21 ??? (fintersects ?) U V).
-
-definition fintersectsS:
- ∀o:BP. Ω^(form o) × Ω^(form o) ⇒_1 Ω^(form o).
- intros (o); constructor 1;
- [ apply (λo: basic_pair.λa,b: Ω^(form o).{c | BPext o c ⊆ BPextS o a ∩ BPextS o b });
- intros; simplify; apply (.= (†e)‡#); apply refl1
- | intros; split; simplify; intros;
- [ apply (. #‡((†e^-1)‡(†e1^-1))); assumption
- | apply (. #‡((†e)‡(†e1))); assumption]]
-qed.
-
-interpretation "fintersectsS" 'fintersects U V = (fun21 ??? (fintersectsS ?) U V).
-
-definition relS: ∀o: BP. (concr o) × Ω^(form o) ⇒_1 CPROP.
- intros (o); constructor 1;
- [ apply (λx:concr o.λS: Ω^(form o).∃y:form o.y ∈ S ∧ x ⊩⎽o y);
- | intros; split; intros; cases e2; exists [1,3: apply w]
- [ apply (. (#‡e1^-1)‡(e^-1‡#)); assumption
- | apply (. (#‡e1)‡(e‡#)); assumption]]
-qed.
-
-interpretation "basic pair relation for subsets" 'Vdash2 x y c = (fun21 (concr ?) ?? (relS c) x y).
-interpretation "basic pair relation for subsets (non applied)" 'Vdash c = (fun21 ??? (relS c)).
-*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/basic_pairs_to_o-basic_pairs.ma".
-include "formal_topology/o-basic_pairs_to_o-basic_topologies.ma".
-include "formal_topology/basic_pairs.ma".
-include "formal_topology/basic_topologies.ma".
-
-definition basic_topology_of_basic_pair: basic_pair → basic_topology.
- intro bp;
- letin obt ≝ (OR (BP_to_OBP bp));
- constructor 1;
- [ apply (form bp);
- | apply (oA obt);
- | apply (oJ obt);
- | apply (oA_is_saturation obt);
- | apply (oJ_is_reduction obt);
- | apply (Ocompatibility obt); ]
-qed.
-
-definition continuous_relation_of_relation_pair:
- ∀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 ocr ≝ (OR⎽⇒ (BP_to_OBP⎽⇒ rp));
- constructor 1;
- [ apply (rp \sub \f);
- | apply (Oreduced ?? ocr);
- | apply (Osaturated ?? ocr); ]
-qed.
-
-alias symbol "compose" (instance 3) = "category1 composition".
-alias symbol "compose" (instance 3) = "category1 composition".
-record functor1 (C1: category1) (C2: category1) : Type2 ≝
- { map_objs1:1> C1 → C2;
- map_arrows1: ∀S,T. unary_morphism1 (arrows1 ? S T) (arrows1 ? (map_objs1 S) (map_objs1 T));
- respects_id1: ∀o:C1. map_arrows1 ?? (id1 ? o) =_1 id1 ? (map_objs1 o);
- respects_comp1:
- ∀o1,o2,o3.∀f1:arrows1 ? o1 o2.∀f2:arrows1 ? o2 o3.
- map_arrows1 ?? (f2 ∘ f1) =_1 map_arrows1 ?? f2 ∘ map_arrows1 ?? f1}.
-
-(*
-definition BTop_of_BP: functor1 BP BTop.
- constructor 1;
- [ apply basic_topology_of_basic_pair
- | intros; constructor 1 [ apply continuous_relation_of_relation_pair; ]
- | simplify; intro;
- ]
-qed.
-
-lemma BBBB_faithful : failthful2 ?? VVV
-FIXME
-*)
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/basic_pairs.ma".
-include "formal_topology/o-basic_pairs.ma".
-include "formal_topology/relations_to_o-algebra.ma".
-
-definition o_basic_pair_of_basic_pair: basic_pair → Obasic_pair.
- intro b;
- constructor 1;
- [ apply (POW (concr b));
- | apply (POW (form b));
- | apply (POW⎽⇒ ?); apply (rel b); ]
-qed.
-
-definition o_relation_pair_of_relation_pair:
- ∀BP1,BP2. relation_pair BP1 BP2 →
- Orelation_pair (o_basic_pair_of_basic_pair BP1) (o_basic_pair_of_basic_pair BP2).
- intros;
- constructor 1;
- [ unfold o_basic_pair_of_basic_pair; simplify; apply (POW⎽⇒ ?); apply (r\sub \c);
- | apply (map_arrows2 ?? POW (form BP1) (form BP2) (r \sub \f));
- | apply (.= (respects_comp2 ?? POW (concr BP1) (concr BP2) (form BP2) r\sub\c (⊩\sub BP2) )^-1);
- cut ( ⊩ \sub BP2 ∘ r \sub \c =_12 r\sub\f ∘ ⊩ \sub BP1) as H;
- [ apply (.= †H);
- apply (respects_comp2 ?? POW (concr BP1) (form BP1) (form BP2) (⊩\sub BP1) r\sub\f);
- | apply commute;]]
-qed.
-
-lemma o_relation_pair_of_relation_pair_is_morphism :
- ∀S,T:category2_of_category1 BP.
- ∀a,b:arrows2 (category2_of_category1 BP) S T.a=b →
- (eq2 (arrows2 OBP (o_basic_pair_of_basic_pair S) (o_basic_pair_of_basic_pair T)))
- (o_relation_pair_of_relation_pair S T a) (o_relation_pair_of_relation_pair S T b).
-intros 2 (S T);
- intros (a b Eab); split; unfold o_relation_pair_of_relation_pair; simplify;
- unfold o_basic_pair_of_basic_pair; simplify;
- [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x);
- | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
- | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
- | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
- simplify;
- apply (prop12);
- apply (.= (respects_comp2 ?? POW (concr S) (concr T) (form T) (a\sub\c) (⊩\sub T))^-1);
- apply sym2;
- apply (.= (respects_comp2 ?? POW (concr S) (concr T) (form T) (b\sub\c) (⊩\sub T))^-1);
- apply sym2;
- apply prop12;
- apply Eab;
-qed.
-
-lemma o_relation_pair_of_relation_pair_morphism :
- ∀S,T:category2_of_category1 BP.
- unary_morphism2 (arrows2 (category2_of_category1 BP) S T)
- (arrows2 OBP (o_basic_pair_of_basic_pair S) (o_basic_pair_of_basic_pair T)).
-intros (S T);
- constructor 1;
- [ apply (o_relation_pair_of_relation_pair S T);
- | apply (o_relation_pair_of_relation_pair_is_morphism S T)]
-qed.
-
-lemma o_relation_pair_of_relation_pair_morphism_respects_id:
- ∀o:category2_of_category1 BP.
- o_relation_pair_of_relation_pair_morphism o o (id2 (category2_of_category1 BP) o)
- = id2 OBP (o_basic_pair_of_basic_pair o).
- simplify; intros; whd; split;
- [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x);
- | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
- | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
- | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
- simplify;
- apply prop12;
- apply prop22;[2,4,6,8: apply rule #;]
- apply (respects_id2 ?? POW (concr o));
-qed.
-
-lemma o_relation_pair_of_relation_pair_morphism_respects_comp:
- ∀o1,o2,o3:category2_of_category1 BP.
- ∀f1:arrows2 (category2_of_category1 BP) o1 o2.
- ∀f2:arrows2 (category2_of_category1 BP) o2 o3.
- (eq2 (arrows2 OBP (o_basic_pair_of_basic_pair o1) (o_basic_pair_of_basic_pair o3)))
- (o_relation_pair_of_relation_pair_morphism o1 o3 (f2 ∘ f1))
- (comp2 OBP ???
- (o_relation_pair_of_relation_pair_morphism o1 o2 f1)
- (o_relation_pair_of_relation_pair_morphism o2 o3 f2)).
- simplify; intros; whd; split;
- [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x);
- | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
- | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
- | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
- simplify;
- apply prop12;
- apply prop22;[2,4,6,8: apply rule #;]
- apply (respects_comp2 ?? POW (concr o1) (concr o2) (concr o3) f1\sub\c f2\sub\c);
-qed.
-
-definition BP_to_OBP: carr3 (arrows3 CAT2 (category2_of_category1 BP) OBP).
- constructor 1;
- [ apply o_basic_pair_of_basic_pair;
- | intros; apply o_relation_pair_of_relation_pair_morphism;
- | apply o_relation_pair_of_relation_pair_morphism_respects_id;
- | apply o_relation_pair_of_relation_pair_morphism_respects_comp;]
-qed.
-
-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 (.= respects_comp2 ?? POW (concr S) (concr T) (form T) g \sub \c (⊩ \sub T));
- apply sym2;
- apply e;
-qed.
-
-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[
- constructor 1; [apply gc|apply gf]
- apply (POW_faithful);
- apply (let xxxx ≝POW in .= respects_comp2 ?? POW (concr S) (concr T) (form T) gc (rel T));
- apply rule (.= Hgc‡#);
- apply (.= Ocommute ?? f);
- apply (.= #‡Hgf^-1);
- apply (let xxxx ≝POW in (respects_comp2 ?? POW (concr S) (form S) (form T) (rel S) gf)^-1)]
- split;
- [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x);
- | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
- | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
- | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
- simplify; apply (†(Hgc‡#));
-qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/relations.ma".
-include "formal_topology/saturations.ma".
-
-record basic_topology: Type1 ≝
- { carrbt:> REL;
- A: Ω^carrbt ⇒_1 Ω^carrbt;
- J: Ω^carrbt ⇒_1 Ω^carrbt;
- A_is_saturation: is_saturation ? A;
- J_is_reduction: is_reduction ? J;
- compatibility: ∀U,V. (A U ≬ J V) =_1 (U ≬ J V)
- }.
-
-record continuous_relation (S,T: basic_topology) : Type1 ≝
- { cont_rel:> S ⇒_\r1 T;
- reduced: ∀U. U =_1 J ? U → image_coercion ?? cont_rel U =_1 J ? (image_coercion ?? cont_rel U);
- saturated: ∀U. U =_1 A ? U → (cont_rel)⎻* U = _1A ? ((cont_rel)⎻* U)
- }.
-
-definition continuous_relation_setoid: basic_topology → basic_topology → setoid1.
- intros (S T); constructor 1;
- [ apply (continuous_relation S T)
- | constructor 1;
- [ apply (λr,s:continuous_relation S T.∀b. A ? (ext ?? r b) = A ? (ext ?? s b));
- | simplify; intros; apply refl1;
- | simplify; intros (x y H); apply sym1; apply H
- | simplify; intros; apply trans1; [2: apply f |3: apply f1; |1: skip]]]
-qed.
-
-definition continuos_relation_of_continuous_relation_setoid :
- ∀P,Q. continuous_relation_setoid P Q → continuous_relation P Q ≝ λP,Q,x.x.
-coercion continuos_relation_of_continuous_relation_setoid.
-
-axiom continuous_relation_eq':
- ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
- a = a' → ∀X.a⎻* (A o1 X) = a'⎻* (A o1 X).
-(*
- intros; split; intro; unfold minus_star_image; simplify; intros;
- [ cut (ext ?? a a1 ⊆ A ? X); [2: intros 2; apply (H1 a2); cases f1; assumption;]
- lapply (if ?? (A_is_saturation ???) Hcut); clear Hcut;
- cut (A ? (ext ?? a' a1) ⊆ A ? X); [2: apply (. (H ?)‡#); assumption]
- lapply (fi ?? (A_is_saturation ???) Hcut);
- apply (Hletin1 x); change with (x ∈ ext ?? a' a1); split; simplify;
- [ apply I | assumption ]
- | cut (ext ?? a' a1 ⊆ A ? X); [2: intros 2; apply (H1 a2); cases f1; assumption;]
- lapply (if ?? (A_is_saturation ???) Hcut); clear Hcut;
- cut (A ? (ext ?? a a1) ⊆ A ? X); [2: apply (. (H ?)\sup -1‡#); assumption]
- lapply (fi ?? (A_is_saturation ???) Hcut);
- apply (Hletin1 x); change with (x ∈ ext ?? a a1); split; simplify;
- [ apply I | assumption ]]
-qed.*)
-
-lemma continuous_relation_eq_inv':
- ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
- (∀X.a⎻* (A o1 X) = a'⎻* (A o1 X)) → a=a'.
- intros 6;
- cut (∀a,a': continuous_relation_setoid o1 o2.
- (∀X.a⎻* (A o1 X) = a'⎻* (A o1 X)) →
- ∀V:o2. A ? (ext ?? a' V) ⊆ A ? (ext ?? a V));
- [2: clear b f a' a; intros;
- lapply depth=0 (λV.saturation_expansive ??? (extS ?? a V)); [2: apply A_is_saturation;|skip]
- (* fundamental adjunction here! to be taken out *)
- cut (∀V:Ω^o2.V ⊆ a⎻* (A ? (extS ?? a V)));
- [2: intro; intros 2; unfold minus_star_image; simplify; intros;
- apply (Hletin V1 x); whd; split; [ exact I | exists; [apply a1] split; assumption]]
- clear Hletin;
- cut (∀V:Ω^o2.V ⊆ a'⎻* (A ? (extS ?? a V)));
- [2: intro; apply (. #‡(f ?)^-1); apply Hcut] clear f Hcut;
- (* second half of the fundamental adjunction here! to be taken out too *)
- intro; lapply (Hcut1 {(V)}); clear Hcut1;
- unfold minus_star_image in Hletin; unfold singleton in Hletin; simplify in Hletin;
- whd in Hletin; whd in Hletin:(?→?→%); simplify in Hletin;
- apply (if ?? (A_is_saturation ???));
- intros 2 (x H); lapply (Hletin V ? x ?);
- [ apply refl | unfold foo; apply H; ]
- change with (x ∈ A ? (ext ?? a V));
- apply (. #‡(†(extS_singleton ????)^-1));
- assumption;]
- split; apply Hcut; [2: assumption | intro; apply sym1; apply f]
-qed.
-
-definition continuous_relation_comp:
- ∀o1,o2,o3.
- continuous_relation_setoid o1 o2 →
- continuous_relation_setoid o2 o3 →
- continuous_relation_setoid o1 o3.
- intros (o1 o2 o3 r s); constructor 1;
- [ alias symbol "compose" (instance 1) = "category1 composition".
-apply (s ∘ r)
- | intros;
- apply sym1;
- (*change in ⊢ (? ? ? (? ? ? ? %) ?) with (image_coercion ?? (s ∘ r) U);*)
- apply (.= †(image_comp ??????));
- apply (.= (reduced ?? s (image_coercion ?? r U) ?)^-1);
- [ apply (.= (reduced ?????)); [ assumption | apply refl1 ]
- | change in ⊢ (? ? ? % ?) with ((image_coercion ?? s ∘ image_coercion ?? r) U);
- apply (.= (image_comp ??????)^-1);
- apply refl1]
- | intros;
- apply sym1;
- apply (.= †(minus_star_image_comp ??? s r ?));
- apply (.= (saturated ?? s ((r)⎻* U) ?)^-1);
- [ apply (.= (saturated ?????)); [ assumption | apply refl1 ]
- | change in ⊢ (? ? ? % ?) with ((s⎻* ∘ r⎻* ) U);
- apply (.= (minus_star_image_comp ??????)^-1);
- apply refl1]]
-qed.
-
-definition BTop: category1.
- constructor 1;
- [ apply basic_topology
- | apply continuous_relation_setoid
- | intro; constructor 1;
- [ apply id1
- | intros;
- apply (.= (image_id ??));
- apply sym1;
- apply (.= †(image_id ??));
- apply sym1;
- assumption
- | intros;
- apply (.= (minus_star_image_id ??));
- apply sym1;
- apply (.= †(minus_star_image_id ??));
- apply sym1;
- assumption]
- | intros; constructor 1;
- [ apply continuous_relation_comp;
- | intros; simplify; intro x; simplify;
- lapply depth=0 (continuous_relation_eq' ???? e) as H';
- lapply depth=0 (continuous_relation_eq' ???? e1) as H1';
- letin K ≝ (λX.H1' ((a)⎻* (A ? X))); clearbody K;
- cut (∀X:Ω \sup o1.
- (b)⎻* (A o2 ((a)⎻* (A o1 X)))
- =_1 (b')⎻* (A o2 ((a')⎻* (A o1 X))));
- [2: intro; apply sym1;
- apply (.= (†(†((H' X)^-1)))); apply sym1; apply (K X);]
- clear K H' H1';
-alias symbol "powerset" (instance 5) = "powerset low".
-alias symbol "compose" (instance 2) = "category1 composition".
-cut (∀X:Ω^o1.
- ((b ∘ a))⎻* (A o1 X) =_1 ((b'∘a'))⎻* (A o1 X));
- [2: intro; unfold foo;
- apply (.= (minus_star_image_comp ??????));
- change in ⊢ (? ? ? % ?) with ((b)⎻* ((a)⎻* (A o1 X)));
- apply (.= †(saturated ?????));
- [ apply ((saturation_idempotent ????)^-1); apply A_is_saturation ]
- apply sym1;
- apply (.= (minus_star_image_comp ??????));
- change in ⊢ (? ? ? % ?) with ((b')⎻* ((a')⎻* (A o1 X)));
- apply (.= †(saturated ?????));
- [ apply ((saturation_idempotent ????)^-1); apply A_is_saturation ]
- apply ((Hcut X)^-1)]
- clear Hcut; generalize in match x; clear x;
- apply (continuous_relation_eq_inv');
- apply Hcut1;]
- | intros; simplify; intro; do 2 (unfold continuous_relation_comp); simplify;
- alias symbol "trans" (instance 1) = "trans1".
-alias symbol "refl" (instance 5) = "refl1".
-alias symbol "prop2" (instance 3) = "prop21".
-alias symbol "prop1" (instance 2) = "prop11".
-alias symbol "assoc" (instance 4) = "category1 assoc".
-apply (.= †(ASSOC‡#));
- apply refl1
- | intros; simplify; intro; unfold continuous_relation_comp; simplify;
- apply (.= †((id_neutral_right1 ????)‡#));
- apply refl1
- | intros; simplify; intro; simplify;
- apply (.= †((id_neutral_left1 ????)‡#));
- apply refl1]
-qed.
-
-(*
-(*CSC: unused! *)
-(* this proof is more logic-oriented than set/lattice oriented *)
-theorem continuous_relation_eqS:
- ∀o1,o2:basic_topology.∀a,a': continuous_relation_setoid o1 o2.
- a = a' → ∀X. A ? (extS ?? a X) = A ? (extS ?? a' X).
- intros;
- cut (∀a: arrows1 ? o1 ?.∀x. x ∈ extS ?? a X → ∃y:o2.y ∈ X ∧ x ∈ ext ?? a y);
- [2: intros; cases f; clear f; cases H1; exists [apply w] cases x1; split;
- try assumption; split; assumption]
- cut (∀a,a':continuous_relation_setoid o1 o2.eq1 ? a a' → ∀x. x ∈ extS ?? a X → ∃y:o2. y ∈ X ∧ x ∈ A ? (ext ?? a' y));
- [2: intros; cases (Hcut ?? f); exists; [apply w] cases x1; split; try assumption;
- apply (. #‡(H1 ?));
- apply (saturation_expansive ?? (A_is_saturation o1) (ext ?? a1 w) x);
- assumption;] clear Hcut;
- split; apply (if ?? (A_is_saturation ???)); intros 2;
- [lapply (Hcut1 a a' H a1 f) | lapply (Hcut1 a' a (H \sup -1) a1 f)]
- cases Hletin; clear Hletin; cases x; clear x;
- cut (∀a: arrows1 ? o1 ?. ext ?? a w ⊆ extS ?? a X);
- [2,4: intros 3; cases f3; clear f3; simplify in f5; split; try assumption;
- exists [1,3: apply w] split; assumption;]
- cut (∀a. A ? (ext o1 o2 a w) ⊆ A ? (extS o1 o2 a X));
- [2,4: intros; apply saturation_monotone; try (apply A_is_saturation); apply Hcut;]
- apply Hcut2; assumption.
-qed.
-*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/basic_topologies.ma".
-include "formal_topology/o-basic_topologies.ma".
-include "formal_topology/relations_to_o-algebra.ma".
-
-definition o_basic_topology_of_basic_topology: basic_topology → Obasic_topology.
- intros (b); constructor 1;
- [ apply (POW' b) | apply (A b) | apply (J b);
- | apply (A_is_saturation b) | apply (J_is_reduction b) | apply (compatibility b) ]
-qed.
-
-definition o_continuous_relation_of_continuous_relation:
- ∀BT1,BT2.continuous_relation BT1 BT2 →
- Ocontinuous_relation (o_basic_topology_of_basic_topology BT1) (o_basic_topology_of_basic_topology BT2).
- intros (BT1 BT2 c); constructor 1;
- [ apply (orelation_of_relation ?? c) | apply (reduced ?? c) | apply (saturated ?? c) ]
-qed.
-
-axiom daemon: False.
-
-lemma o_continuous_relation_of_continuous_relation_morphism :
- ∀S,T:category2_of_category1 BTop.
- unary_morphism2 (arrows2 (category2_of_category1 BTop) S T)
- (arrows2 OBTop (o_basic_topology_of_basic_topology S) (o_basic_topology_of_basic_topology T)).
-intros (S T);
- constructor 1;
- [ apply (o_continuous_relation_of_continuous_relation S T);
- | cases daemon (*apply (o_relation_pair_of_relation_pair_is_morphism S T)*)]
-qed.
-
-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;
- | cases daemon (*apply o_relation_topology_of_relation_topology_morphism_respects_id*);
- | cases daemon (*apply o_relation_topology_of_relation_topology_morphism_respects_comp*);]
-qed.
-
-theorem BTop_to_OBTop_faithful: faithful2 ?? BTop_to_OBTop.
- intros 5; apply (continuous_relation_eq_inv' o1 o2 f g); apply e;
-qed.
-
-include "formal_topology/notation.ma".
-
-theorem BTop_to_OBTop_full: full2 ?? BTop_to_OBTop.
- intros 3 (S T);
- cases (POW_full (carrbt S) (carrbt T) (Ocont_rel ?? f)) (g Hg);
- (* cases Hg; *)
- exists [
- constructor 1;
- [ apply g
- | unfold image_coercion; cases daemon (*apply hide; intros; lapply (Oreduced ?? f ? e); unfold image_coercion;
- cases Hg; lapply (e3 U) as K; apply (.= K);
- apply (.= Hletin); apply rule (†(K^-1)); *)
- | cases daemon (* apply hide; intros; lapply (Osaturated ?? f ? e);
- cases Hg; lapply (e1 U) as K; apply (.= K);
- apply (.= Hletin); apply rule (†(K^-1)); *)
- ]
- | simplify; unfold BTop_to_OBTop; simplify;
- cases Hg; unfold o_continuous_relation_of_continuous_relation_morphism;
- simplify;
- change with ((orelation_of_relation ?? g)⎻* ∘ oA (o_basic_topology_of_basic_topology S) =
- f⎻* ∘ oA (o_basic_topology_of_basic_topology S));
-
-
- change with (g⎻* ∘ oA (o_basic_topology_of_basic_topology S) =
- f⎻* ∘ oA (o_basic_topology_of_basic_topology S));
- apply sym2; whd in T;
- intro;
- apply trans2; [2: apply sym2; [2: apply Hg;
-
- whd in ⊢ (?(??%%)???);
- apply (.= Hg^-1);
- unfold o_continuous_relation_of_continuous_relation_morphism; simplify;
- intro; simplify;
- unfold image_coercion; cases Hg; whd; simplify; intro; simplify;
-qed.
-*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/cprop_connectives.ma".
-
-notation "hvbox(a break = \sub \ID b)" non associative with precedence 45
-for @{ 'eqID $a $b }.
-
-notation > "hvbox(a break =_\ID b)" non associative with precedence 45
-for @{ cic:/matita/logic/equality/eq.ind#xpointer(1/1) ? $a $b }.
-
-interpretation "ID eq" 'eqID x y = (cic:/matita/logic/equality/eq.ind#xpointer(1/1) ? x y).
-
-record equivalence_relation (A:Type0) : Type1 ≝
- { eq_rel:2> A → A → CProp0;
- refl: reflexive ? eq_rel;
- sym: symmetric ? eq_rel;
- trans: transitive ? eq_rel
- }.
-
-record setoid : Type1 ≝
- { carr:> Type0;
- eq: equivalence_relation carr
- }.
-
-record equivalence_relation1 (A:Type1) : Type2 ≝
- { eq_rel1:2> A → A → CProp1;
- refl1: reflexive1 ? eq_rel1;
- sym1: symmetric1 ? eq_rel1;
- trans1: transitive1 ? eq_rel1
- }.
-
-record setoid1: Type2 ≝
- { carr1:> Type1;
- eq1: equivalence_relation1 carr1
- }.
-
-definition setoid1_of_setoid: setoid → setoid1.
- intro;
- constructor 1;
- [ apply (carr s)
- | constructor 1;
- [ apply (eq_rel s);
- apply (eq s)
- | apply (refl s)
- | apply (sym s)
- | apply (trans s)]]
-qed.
-
-coercion setoid1_of_setoid.
-prefer coercion Type_OF_setoid.
-
-record equivalence_relation2 (A:Type2) : Type3 ≝
- { eq_rel2:2> A → A → CProp2;
- refl2: reflexive2 ? eq_rel2;
- sym2: symmetric2 ? eq_rel2;
- trans2: transitive2 ? eq_rel2
- }.
-
-record setoid2: Type3 ≝
- { carr2:> Type2;
- eq2: equivalence_relation2 carr2
- }.
-
-definition setoid2_of_setoid1: setoid1 → setoid2.
- intro;
- constructor 1;
- [ apply (carr1 s)
- | constructor 1;
- [ apply (eq_rel1 s);
- apply (eq1 s)
- | apply (refl1 s)
- | apply (sym1 s)
- | apply (trans1 s)]]
-qed.
-
-coercion setoid2_of_setoid1.
-prefer coercion Type_OF_setoid2.
-prefer coercion Type_OF_setoid.
-prefer coercion Type_OF_setoid1.
-(* we prefer 0 < 1 < 2 *)
-
-record equivalence_relation3 (A:Type3) : Type4 ≝
- { eq_rel3:2> A → A → CProp3;
- refl3: reflexive3 ? eq_rel3;
- sym3: symmetric3 ? eq_rel3;
- trans3: transitive3 ? eq_rel3
- }.
-
-record setoid3: Type4 ≝
- { carr3:> Type3;
- eq3: equivalence_relation3 carr3
- }.
-
-interpretation "setoid3 eq" 'eq t x y = (eq_rel3 ? (eq3 t) x y).
-interpretation "setoid2 eq" 'eq t x y = (eq_rel2 ? (eq2 t) x y).
-interpretation "setoid1 eq" 'eq t x y = (eq_rel1 ? (eq1 t) x y).
-interpretation "setoid eq" 'eq t x y = (eq_rel ? (eq t) x y).
-
-notation > "hvbox(a break =_12 b)" non associative with precedence 45
-for @{ eq_rel2 (carr2 (setoid2_of_setoid1 ?)) (eq2 (setoid2_of_setoid1 ?)) $a $b }.
-notation > "hvbox(a break =_0 b)" non associative with precedence 45
-for @{ eq_rel ? (eq ?) $a $b }.
-notation > "hvbox(a break =_1 b)" non associative with precedence 45
-for @{ eq_rel1 ? (eq1 ?) $a $b }.
-notation > "hvbox(a break =_2 b)" non associative with precedence 45
-for @{ eq_rel2 ? (eq2 ?) $a $b }.
-notation > "hvbox(a break =_3 b)" non associative with precedence 45
-for @{ eq_rel3 ? (eq3 ?) $a $b }.
-
-interpretation "setoid3 symmetry" 'invert r = (sym3 ???? r).
-interpretation "setoid2 symmetry" 'invert r = (sym2 ???? r).
-interpretation "setoid1 symmetry" 'invert r = (sym1 ???? r).
-interpretation "setoid symmetry" 'invert r = (sym ???? r).
-notation ".= r" with precedence 50 for @{'trans $r}.
-interpretation "trans3" 'trans r = (trans3 ????? r).
-interpretation "trans2" 'trans r = (trans2 ????? r).
-interpretation "trans1" 'trans r = (trans1 ????? r).
-interpretation "trans" 'trans r = (trans ????? r).
-
-record unary_morphism (A,B: setoid) : Type0 ≝
- { fun1:1> A → B;
- prop1: ∀a,a'. eq ? a a' → eq ? (fun1 a) (fun1 a')
- }.
-
-record unary_morphism1 (A,B: setoid1) : Type1 ≝
- { fun11:1> A → B;
- prop11: ∀a,a'. eq1 ? a a' → eq1 ? (fun11 a) (fun11 a')
- }.
-
-record unary_morphism2 (A,B: setoid2) : Type2 ≝
- { fun12:1> A → B;
- prop12: ∀a,a'. eq2 ? a a' → eq2 ? (fun12 a) (fun12 a')
- }.
-
-record unary_morphism3 (A,B: setoid3) : Type3 ≝
- { fun13:1> A → B;
- prop13: ∀a,a'. eq3 ? a a' → eq3 ? (fun13 a) (fun13 a')
- }.
-
-record binary_morphism (A,B,C:setoid) : Type0 ≝
- { fun2:2> A → B → C;
- prop2: ∀a,a',b,b'. eq ? a a' → eq ? b b' → eq ? (fun2 a b) (fun2 a' b')
- }.
-
-record binary_morphism1 (A,B,C:setoid1) : Type1 ≝
- { fun21:2> A → B → C;
- prop21: ∀a,a',b,b'. eq1 ? a a' → eq1 ? b b' → eq1 ? (fun21 a b) (fun21 a' b')
- }.
-
-record binary_morphism2 (A,B,C:setoid2) : Type2 ≝
- { fun22:2> A → B → C;
- prop22: ∀a,a',b,b'. eq2 ? a a' → eq2 ? b b' → eq2 ? (fun22 a b) (fun22 a' b')
- }.
-
-record binary_morphism3 (A,B,C:setoid3) : Type3 ≝
- { fun23:2> A → B → C;
- prop23: ∀a,a',b,b'. eq3 ? a a' → eq3 ? b b' → eq3 ? (fun23 a b) (fun23 a' b')
- }.
-
-notation "† c" with precedence 90 for @{'prop1 $c }.
-notation "l ‡ r" with precedence 90 for @{'prop2 $l $r }.
-notation "#" with precedence 90 for @{'refl}.
-interpretation "prop1" 'prop1 c = (prop1 ????? c).
-interpretation "prop11" 'prop1 c = (prop11 ????? c).
-interpretation "prop12" 'prop1 c = (prop12 ????? c).
-interpretation "prop13" 'prop1 c = (prop13 ????? c).
-interpretation "prop2" 'prop2 l r = (prop2 ???????? l r).
-interpretation "prop21" 'prop2 l r = (prop21 ???????? l r).
-interpretation "prop22" 'prop2 l r = (prop22 ???????? l r).
-interpretation "prop23" 'prop2 l r = (prop23 ???????? l r).
-interpretation "refl" 'refl = (refl ???).
-interpretation "refl1" 'refl = (refl1 ???).
-interpretation "refl2" 'refl = (refl2 ???).
-interpretation "refl3" 'refl = (refl3 ???).
-
-notation > "A × term 74 B ⇒ term 19 C" non associative with precedence 72 for @{'binary_morphism0 $A $B $C}.
-notation > "A × term 74 B ⇒_1 term 19 C" non associative with precedence 72 for @{'binary_morphism1 $A $B $C}.
-notation > "A × term 74 B ⇒_2 term 19 C" non associative with precedence 72 for @{'binary_morphism2 $A $B $C}.
-notation > "A × term 74 B ⇒_3 term 19 C" non associative with precedence 72 for @{'binary_morphism3 $A $B $C}.
-notation > "B ⇒_1 C" right associative with precedence 72 for @{'arrows1_SET $B $C}.
-notation > "B ⇒_1. C" right associative with precedence 72 for @{'arrows1_SETlow $B $C}.
-notation > "B ⇒_2 C" right associative with precedence 72 for @{'arrows2_SET1 $B $C}.
-notation > "B ⇒_2. C" right associative with precedence 72 for @{'arrows2_SET1low $B $C}.
-
-notation "A × term 74 B ⇒ term 19 C" non associative with precedence 72 for @{'binary_morphism0 $A $B $C}.
-notation "A × term 74 B ⇒\sub 1 term 19 C" non associative with precedence 72 for @{'binary_morphism1 $A $B $C}.
-notation "A × term 74 B ⇒\sub 2 term 19 C" non associative with precedence 72 for @{'binary_morphism2 $A $B $C}.
-notation "A × term 74 B ⇒\sub 3 term 19 C" non associative with precedence 72 for @{'binary_morphism3 $A $B $C}.
-notation "B ⇒\sub 1 C" right associative with precedence 72 for @{'arrows1_SET $B $C}.
-notation "B ⇒\sub 2 C" right associative with precedence 72 for @{'arrows2_SET1 $B $C}.
-notation "B ⇒\sub 1. C" right associative with precedence 72 for @{'arrows1_SETlow $B $C}.
-notation "B ⇒\sub 2. C" right associative with precedence 72 for @{'arrows2_SET1low $B $C}.
-
-interpretation "'binary_morphism0" 'binary_morphism0 A B C = (binary_morphism A B C).
-interpretation "'arrows2_SET1 low" 'arrows2_SET1 A B = (unary_morphism2 A B).
-interpretation "'arrows2_SET1low" 'arrows2_SET1low A B = (unary_morphism2 A B).
-interpretation "'binary_morphism1" 'binary_morphism1 A B C = (binary_morphism1 A B C).
-interpretation "'binary_morphism2" 'binary_morphism2 A B C = (binary_morphism2 A B C).
-interpretation "'binary_morphism3" 'binary_morphism3 A B C = (binary_morphism3 A B C).
-interpretation "'arrows1_SET low" 'arrows1_SET A B = (unary_morphism1 A B).
-interpretation "'arrows1_SETlow" 'arrows1_SETlow A B = (unary_morphism1 A B).
-
-
-definition unary_morphism2_of_unary_morphism1:
- ∀S,T:setoid1.unary_morphism1 S T → unary_morphism2 (setoid2_of_setoid1 S) T.
- intros;
- constructor 1;
- [ apply (fun11 ?? u);
- | apply (prop11 ?? u); ]
-qed.
-
-definition CPROP: setoid1.
- constructor 1;
- [ apply CProp0
- | constructor 1;
- [ apply Iff
- | intros 1; split; intro; assumption
- | intros 3; cases i; split; assumption
- | intros 5; cases i; cases i1; split; intro;
- [ apply (f2 (f x1)) | apply (f1 (f3 z1))]]]
-qed.
-
-definition CProp0_of_CPROP: carr1 CPROP → CProp0 ≝ λx.x.
-coercion CProp0_of_CPROP.
-
-alias symbol "eq" = "setoid1 eq".
-definition fi': ∀A,B:CPROP. A = B → B → A.
- intros; apply (fi ?? e); assumption.
-qed.
-
-notation ". r" with precedence 50 for @{'fi $r}.
-interpretation "fi" 'fi r = (fi' ?? r).
-
-definition and_morphism: binary_morphism1 CPROP CPROP CPROP.
- constructor 1;
- [ apply And
- | intros; split; intro; cases a1; split;
- [ apply (if ?? e a2)
- | apply (if ?? e1 b1)
- | apply (fi ?? e a2)
- | apply (fi ?? e1 b1)]]
-qed.
-
-interpretation "and_morphism" 'and a b = (fun21 ??? and_morphism a b).
-
-definition or_morphism: binary_morphism1 CPROP CPROP CPROP.
- constructor 1;
- [ apply Or
- | intros; split; intro; cases o; [1,3:left |2,4: right]
- [ apply (if ?? e a1)
- | apply (fi ?? e a1)
- | apply (if ?? e1 b1)
- | apply (fi ?? e1 b1)]]
-qed.
-
-interpretation "or_morphism" 'or a b = (fun21 ??? or_morphism a b).
-
-definition if_morphism: binary_morphism1 CPROP CPROP CPROP.
- constructor 1;
- [ apply (λA,B. A → B)
- | intros; split; intros;
- [ apply (if ?? e1); apply f; apply (fi ?? e); assumption
- | apply (fi ?? e1); apply f; apply (if ?? e); assumption]]
-qed.
-
-notation > "hvbox(a break ∘ b)" left associative with precedence 55 for @{ comp ??? $a $b }.
-record category : Type1 ≝ {
- objs:> Type0;
- arrows: objs → objs → setoid;
- id: ∀o:objs. arrows o o;
- comp: ∀o1,o2,o3. (arrows o1 o2) × (arrows o2 o3) ⇒ (arrows o1 o3);
- comp_assoc: ∀o1,o2,o3,o4.∀a12:arrows o1 ?.∀a23:arrows o2 ?.∀a34:arrows o3 o4.
- (a12 ∘ a23) ∘ a34 =_0 a12 ∘ (a23 ∘ a34);
- id_neutral_left : ∀o1,o2. ∀a: arrows o1 o2. (id o1) ∘ a =_0 a;
- id_neutral_right: ∀o1,o2. ∀a: arrows o1 o2. a ∘ (id o2) =_0 a
-}.
-notation "hvbox(a break ∘ b)" left associative with precedence 55 for @{ 'compose $a $b }.
-
-record category1 : Type2 ≝
- { objs1:> Type1;
- arrows1: objs1 → objs1 → setoid1;
- id1: ∀o:objs1. arrows1 o o;
- comp1: ∀o1,o2,o3. binary_morphism1 (arrows1 o1 o2) (arrows1 o2 o3) (arrows1 o1 o3);
- comp_assoc1: ∀o1,o2,o3,o4. ∀a12,a23,a34.
- comp1 o1 o3 o4 (comp1 o1 o2 o3 a12 a23) a34 =_1 comp1 o1 o2 o4 a12 (comp1 o2 o3 o4 a23 a34);
- id_neutral_right1: ∀o1,o2. ∀a: arrows1 o1 o2. comp1 ??? (id1 o1) a =_1 a;
- id_neutral_left1: ∀o1,o2. ∀a: arrows1 o1 o2. comp1 ??? a (id1 o2) =_1 a
- }.
-
-record category2 : Type3 ≝
- { objs2:> Type2;
- arrows2: objs2 → objs2 → setoid2;
- id2: ∀o:objs2. arrows2 o o;
- comp2: ∀o1,o2,o3. binary_morphism2 (arrows2 o1 o2) (arrows2 o2 o3) (arrows2 o1 o3);
- comp_assoc2: ∀o1,o2,o3,o4. ∀a12,a23,a34.
- comp2 o1 o3 o4 (comp2 o1 o2 o3 a12 a23) a34 =_2 comp2 o1 o2 o4 a12 (comp2 o2 o3 o4 a23 a34);
- id_neutral_right2: ∀o1,o2. ∀a: arrows2 o1 o2. comp2 ??? (id2 o1) a =_2 a;
- id_neutral_left2: ∀o1,o2. ∀a: arrows2 o1 o2. comp2 ??? a (id2 o2) =_2 a
- }.
-
-record category3 : Type4 ≝
- { objs3:> Type3;
- arrows3: objs3 → objs3 → setoid3;
- id3: ∀o:objs3. arrows3 o o;
- comp3: ∀o1,o2,o3. binary_morphism3 (arrows3 o1 o2) (arrows3 o2 o3) (arrows3 o1 o3);
- comp_assoc3: ∀o1,o2,o3,o4. ∀a12,a23,a34.
- comp3 o1 o3 o4 (comp3 o1 o2 o3 a12 a23) a34 =_3 comp3 o1 o2 o4 a12 (comp3 o2 o3 o4 a23 a34);
- id_neutral_right3: ∀o1,o2. ∀a: arrows3 o1 o2. comp3 ??? (id3 o1) a =_3 a;
- id_neutral_left3: ∀o1,o2. ∀a: arrows3 o1 o2. comp3 ??? a (id3 o2) =_3 a
- }.
-
-notation "'ASSOC'" with precedence 90 for @{'assoc}.
-
-interpretation "category2 composition" 'compose x y = (fun22 ??? (comp2 ????) y x).
-interpretation "category2 assoc" 'assoc = (comp_assoc2 ????????).
-interpretation "category1 composition" 'compose x y = (fun21 ??? (comp1 ????) y x).
-interpretation "category1 assoc" 'assoc = (comp_assoc1 ????????).
-interpretation "category composition" 'compose x y = (fun2 ??? (comp ????) y x).
-interpretation "category assoc" 'assoc = (comp_assoc ????????).
-
-definition category2_of_category1: category1 → category2.
- intro;
- constructor 1;
- [ apply (objs1 c);
- | intros; apply (setoid2_of_setoid1 (arrows1 c o o1));
- | apply (id1 c);
- | intros;
- constructor 1;
- [ intros; apply (comp1 c o1 o2 o3 c1 c2);
- | intros; unfold setoid2_of_setoid1 in e e1 a a' b b'; simplify in e e1 a a' b b';
- change with ((b∘a) =_1 (b'∘a')); apply (e‡e1); ]
- | intros; simplify; whd in a12 a23 a34; whd; apply rule (ASSOC);
- | intros; simplify; whd in a; whd; apply id_neutral_right1;
- | intros; simplify; whd in a; whd; apply id_neutral_left1; ]
-qed.
-(*coercion category2_of_category1.*)
-
-record functor2 (C1: category2) (C2: category2) : Type3 ≝
- { map_objs2:1> C1 → C2;
- map_arrows2: ∀S,T. unary_morphism2 (arrows2 ? S T) (arrows2 ? (map_objs2 S) (map_objs2 T));
- respects_id2: ∀o:C1. map_arrows2 ?? (id2 ? o) = id2 ? (map_objs2 o);
- respects_comp2:
- ∀o1,o2,o3.∀f1:arrows2 ? o1 o2.∀f2:arrows2 ? o2 o3.
- map_arrows2 ?? (f2 ∘ f1) = map_arrows2 ?? f2 ∘ map_arrows2 ?? f1}.
-
-notation > "F⎽⇒ x" left associative with precedence 60 for @{'map_arrows2 $F $x}.
-notation "F\sub⇒ x" left associative with precedence 60 for @{'map_arrows2 $F $x}.
-interpretation "map_arrows2" 'map_arrows2 F x = (fun12 ?? (map_arrows2 ?? F ??) x).
-
-definition functor2_setoid: category2 → category2 → setoid3.
- intros (C1 C2);
- constructor 1;
- [ apply (functor2 C1 C2);
- | constructor 1;
- [ intros (f g);
- apply (∀c:C1. cic:/matita/logic/equality/eq.ind#xpointer(1/1) ? (f c) (g c));
- | simplify; intros; apply cic:/matita/logic/equality/eq.ind#xpointer(1/1/1);
- | simplify; intros; apply cic:/matita/logic/equality/sym_eq.con; apply H;
- | simplify; intros; apply cic:/matita/logic/equality/trans_eq.con;
- [2: apply H; | skip | apply H1;]]]
-qed.
-
-definition functor2_of_functor2_setoid: ∀S,T. functor2_setoid S T → functor2 S T ≝ λS,T,x.x.
-coercion functor2_of_functor2_setoid.
-
-definition CAT2: category3.
- constructor 1;
- [ apply category2;
- | apply functor2_setoid;
- | intros; constructor 1;
- [ apply (λx.x);
- | intros; constructor 1;
- [ apply (λx.x);
- | intros; assumption;]
- | intros; apply rule #;
- | intros; apply rule #; ]
- | intros; constructor 1;
- [ intros; constructor 1;
- [ intros; apply (c1 (c o));
- | intros; constructor 1;
- [ intro; apply (map_arrows2 ?? c1 ?? (map_arrows2 ?? c ?? c2));
- | intros; apply (††e); ]
- | intros; simplify;
- apply (.= †(respects_id2 : ?));
- apply (respects_id2 : ?);
- | intros; simplify;
- apply (.= †(respects_comp2 : ?));
- apply (respects_comp2 : ?); ]
- | intros; intro; simplify;
- apply (cic:/matita/logic/equality/eq_ind.con ????? (e ?));
- apply (cic:/matita/logic/equality/eq_ind.con ????? (e1 ?));
- constructor 1; ]
- | intros; intro; simplify; constructor 1;
- | intros; intro; simplify; constructor 1;
- | intros; intro; simplify; constructor 1; ]
-qed.
-
-definition category2_of_objs3_CAT2: objs3 CAT2 → category2 ≝ λx.x.
-coercion category2_of_objs3_CAT2.
-
-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;
- [ apply (unary_morphism s s1);
- | constructor 1;
- [ intros (f g); apply (∀a:s. eq ? (f a) (g a));
- | intros 1; simplify; intros; apply refl;
- | simplify; intros; apply sym; apply f;
- | simplify; intros; apply trans; [2: apply f; | skip | apply f1]]]
-qed.
-
-definition SET: category1.
- constructor 1;
- [ apply setoid;
- | apply rule (λS,T:setoid.setoid1_of_setoid (unary_morphism_setoid S T));
- | intros; constructor 1; [ apply (λx:carr o.x); | intros; assumption ]
- | intros; constructor 1; [ intros; constructor 1; [ apply (λx. c1 (c x)); | intros;
- apply († (†e));]
- | intros; whd; intros; simplify; whd in H1; whd in H;
- apply trans; [ apply (b (a' a1)); | lapply (prop1 ?? b (a a1) (a' a1));
- [ apply Hletin | apply (e a1); ] | apply e1; ]]
- | intros; whd; intros; simplify; apply refl;
- | intros; simplify; whd; intros; simplify; apply refl;
- | intros; simplify; whd; intros; simplify; apply refl;
- ]
-qed.
-
-definition setoid_of_SET: objs1 SET → setoid ≝ λx.x.
-coercion setoid_of_SET.
-
-definition unary_morphism_setoid_of_arrows1_SET:
- ∀P,Q.arrows1 SET P Q → unary_morphism_setoid P Q ≝ λP,Q,x.x.
-coercion unary_morphism_setoid_of_arrows1_SET.
-
-interpretation "'arrows1_SET" 'arrows1_SET A B = (arrows1 SET A B).
-
-definition unary_morphism1_setoid1: setoid1 → setoid1 → setoid1.
- intros;
- constructor 1;
- [ apply (unary_morphism1 s s1);
- | constructor 1;
- [ intros (f g);
- alias symbol "eq" = "setoid1 eq".
- apply (∀a: carr1 s. f a = g a);
- | intros 1; simplify; intros; apply refl1;
- | simplify; intros; apply sym1; apply f;
- | simplify; intros; apply trans1; [2: apply f; | skip | apply f1]]]
-qed.
-
-definition unary_morphism1_of_unary_morphism1_setoid1 :
- ∀S,T. unary_morphism1_setoid1 S T → unary_morphism1 S T ≝ λP,Q,x.x.
-coercion unary_morphism1_of_unary_morphism1_setoid1.
-
-definition SET1: objs3 CAT2.
- constructor 1;
- [ apply setoid1;
- | apply rule (λS,T.setoid2_of_setoid1 (unary_morphism1_setoid1 S T));
- | intros; constructor 1; [ apply (λx.x); | intros; assumption ]
- | intros; constructor 1; [ intros; constructor 1; [ apply (λx. c1 (c x)); | intros;
- apply († (†e));]
- | intros; whd; intros; simplify; whd in H1; whd in H;
- apply trans1; [ apply (b (a' a1)); | lapply (prop11 ?? b (a a1) (a' a1));
- [ apply Hletin | apply (e a1); ] | apply e1; ]]
- | intros; whd; intros; simplify; apply refl1;
- | intros; simplify; whd; intros; simplify; apply refl1;
- | intros; simplify; whd; intros; simplify; apply refl1;
- ]
-qed.
-
-interpretation "'arrows2_SET1" 'arrows2_SET1 A B = (arrows2 SET1 A B).
-
-definition setoid1_of_SET1: objs2 SET1 → setoid1 ≝ λx.x.
-coercion setoid1_of_SET1.
-
-definition unary_morphism1_setoid1_of_arrows2_SET1:
- ∀P,Q.arrows2 SET1 P Q → unary_morphism1_setoid1 P Q ≝ λP,Q,x.x.
-coercion unary_morphism1_setoid1_of_arrows2_SET1.
-
-variant objs2_of_category1: objs1 SET → objs2 SET1 ≝ setoid1_of_setoid.
-coercion objs2_of_category1.
-
-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.
-
-
-notation "r \sup *" non associative with precedence 90 for @{'OR_f_star $r}.
-notation > "r *" non associative with precedence 90 for @{'OR_f_star $r}.
-
-notation "r \sup (⎻* )" non associative with precedence 90 for @{'OR_f_minus_star $r}.
-notation > "r⎻*" non associative with precedence 90 for @{'OR_f_minus_star $r}.
-
-notation "r \sup ⎻" non associative with precedence 90 for @{'OR_f_minus $r}.
-notation > "r⎻" non associative with precedence 90 for @{'OR_f_minus $r}.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/basic_pairs.ma".
-
-(* carr1 e' necessario perche' ci sega via la coercion per gli oggetti di REL!
- (confondendola con la coercion per gli oggetti di SET
-record concrete_space : Type1 ≝
- { bp:> BP;
- converges: ∀a: carr1 (concr bp).∀U,V: carr1 (form bp). a ⊩ U → a ⊩ V → a ⊩ (U ↓ V);
- all_covered: ∀x: carr1 (concr bp). x ⊩ form bp
- }.
-
-record convergent_relation_pair (CS1,CS2: concrete_space) : Type1 ≝
- { rp:> arrows1 ? CS1 CS2;
- respects_converges:
- ∀b,c.
- minus_image ?? rp \sub\c (BPextS CS2 (b ↓ c)) =
- BPextS CS1 ((minus_image ?? rp \sub\f b) ↓ (minus_image ?? rp \sub\f c));
- respects_all_covered:
- minus_image ?? rp\sub\c (BPextS CS2 (full_subset (form CS2))) = BPextS CS1 (full_subset (form CS1))
- }.
-
-definition convergent_relation_space_setoid: concrete_space → concrete_space → setoid1.
- intros;
- constructor 1;
- [ apply (convergent_relation_pair c c1)
- | constructor 1;
- [ intros;
- apply (relation_pair_equality c c1 c2 c3);
- | intros 1; apply refl1;
- | intros 2; apply sym1;
- | intros 3; apply trans1]]
-qed.
-
-definition convergent_relation_space_composition:
- ∀o1,o2,o3: concrete_space.
- binary_morphism1
- (convergent_relation_space_setoid o1 o2)
- (convergent_relation_space_setoid o2 o3)
- (convergent_relation_space_setoid o1 o3).
- intros; constructor 1;
- [ intros; whd in c c1 ⊢ %;
- constructor 1;
- [ apply (fun1 ??? (comp1 BP ???)); [apply (bp o2) |*: apply rp; assumption]
- | intros;
- change in ⊢ (? ? ? (? ? ? (? ? ? %) ?) ?) with (c1 \sub \c ∘ c \sub \c);
- change in ⊢ (? ? ? ? (? ? ? ? (? ? ? ? ? (? ? ? (? ? ? %) ?) ?)))
- with (c1 \sub \f ∘ c \sub \f);
- change in ⊢ (? ? ? ? (? ? ? ? (? ? ? ? ? ? (? ? ? (? ? ? %) ?))))
- with (c1 \sub \f ∘ c \sub \f);
- apply (.= (extS_com ??????));
- apply (.= (†(respects_converges ?????)));
- apply (.= (respects_converges ?????));
- apply (.= (†(((extS_com ??????) \sup -1)‡(extS_com ??????)\sup -1)));
- apply refl1;
- | change in ⊢ (? ? ? (? ? ? (? ? ? %) ?) ?) with (c1 \sub \c ∘ c \sub \c);
- apply (.= (extS_com ??????));
- apply (.= (†(respects_all_covered ???)));
- apply (.= respects_all_covered ???);
- apply refl1]
- | intros;
- change with (b ∘ a = b' ∘ a');
- change in H with (rp'' ?? a = rp'' ?? a');
- change in H1 with (rp'' ?? b = rp ?? b');
- apply (.= (H‡H1));
- apply refl1]
-qed.
-
-definition CSPA: category1.
- constructor 1;
- [ apply concrete_space
- | apply convergent_relation_space_setoid
- | intro; constructor 1;
- [ apply id1
- | intros;
- unfold id; simplify;
- apply (.= (equalset_extS_id_X_X ??));
- apply (.= (†((equalset_extS_id_X_X ??)\sup -1‡
- (equalset_extS_id_X_X ??)\sup -1)));
- apply refl1;
- | apply (.= (equalset_extS_id_X_X ??));
- apply refl1]
- | apply convergent_relation_space_composition
- | intros; simplify;
- change with (a34 ∘ (a23 ∘ a12) = (a34 ∘ a23) ∘ a12);
- apply (.= ASSOC1);
- apply refl1
- | intros; simplify;
- change with (a ∘ id1 ? o1 = a);
- apply (.= id_neutral_right1 ????);
- apply refl1
- | intros; simplify;
- change with (id1 ? o2 ∘ a = a);
- apply (.= id_neutral_left1 ????);
- apply refl1]
-qed.
-*)
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/concrete_spaces.ma".
-include "formal_topology/o-concrete_spaces.ma".
-include "formal_topology/basic_pairs_to_o-basic_pairs.ma".
-
-(*
-(* Qui, per fare le cose per bene, ci serve la nozione di funtore categorico *)
-definition o_concrete_space_of_concrete_space: cic:/matita/formal_topology/concrete_spaces/concrete_space.ind#xpointer(1/1) → concrete_space.
- intro;
- constructor 1;
- [ apply (o_basic_pair_of_basic_pair (bp c));
- | lapply depth=0 (downarrow c);
- constructor 1;
- [ apply
- |
- apply (o_operator_of_operator);fintersectsS);
- |
- |
- |
- |
- |
- ]
-qed.
-
-definition o_convergent_relation_pair_of_convergent_relation_pair:
- ∀BP1,BP2.cic:/matita/formal_topology/concrete_spaces/convergent_relation_pair.ind#xpointer(1/1) BP1 BP2 →
- convergent_relation_pair (o_concrete_space_of_concrete_space BP1) (o_concrete_space_of_concrete_space BP2).
- intros;
- constructor 1;
- [ apply (orelation_of_relation ?? (r \sub \c));
- | apply (orelation_of_relation ?? (r \sub \f));
- | lapply (commute ?? r);
- lapply (orelation_of_relation_preserves_equality ???? Hletin);
- apply (.= (orelation_of_relation_preserves_composition (concr BP1) ??? (rel BP2)) ^ -1);
- apply (.= (orelation_of_relation_preserves_equality ???? (commute ?? r)));
- apply (orelation_of_relation_preserves_composition ?? (form BP2) (rel BP1) ?); ]
-qed.
-
-*)
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "logic/connectives.ma".
-
-definition Type4 : Type := Type.
-definition Type3 : Type4 := Type.
-definition Type2 : Type3 := Type.
-definition Type1 : Type2 := Type.
-definition Type0 : Type1 := Type.
-
-definition Type_of_Type0: Type0 → Type := λx.x.
-definition Type_of_Type1: Type1 → Type := λx.x.
-definition Type_of_Type2: Type2 → Type := λx.x.
-definition Type_of_Type3: Type3 → Type := λx.x.
-definition Type_of_Type4: Type4 → Type := λx.x.
-coercion Type_of_Type0.
-coercion Type_of_Type1.
-coercion Type_of_Type2.
-coercion Type_of_Type3.
-coercion Type_of_Type4.
-
-definition CProp0 : Type1 := Type0.
-definition CProp1 : Type2 := Type1.
-definition CProp2 : Type3 := Type2.
-definition CProp3 : Type4 := Type3.
-definition CProp_of_CProp0: CProp0 → CProp ≝ λx.x.
-definition CProp_of_CProp1: CProp1 → CProp ≝ λx.x.
-definition CProp_of_CProp2: CProp2 → CProp ≝ λx.x.
-definition CProp_of_CProp3: CProp3 → CProp ≝ λx.x.
-coercion CProp_of_CProp0.
-coercion CProp_of_CProp1.
-coercion CProp_of_CProp2.
-coercion CProp_of_CProp3.
-
-inductive Or (A,B:CProp0) : CProp0 ≝
- | Left : A → Or A B
- | Right : B → Or A B.
-
-interpretation "constructive or" 'or x y = (Or x y).
-
-inductive Or3 (A,B,C:CProp0) : CProp0 ≝
- | Left3 : A → Or3 A B C
- | Middle3 : B → Or3 A B C
- | Right3 : C → Or3 A B C.
-
-interpretation "constructive ternary or" 'or3 x y z= (Or3 x y z).
-
-notation < "hvbox(a break ∨ b break ∨ c)" with precedence 35 for @{'or3 $a $b $c}.
-
-inductive Or4 (A,B,C,D:CProp0) : CProp0 ≝
- | Left3 : A → Or4 A B C D
- | Middle3 : B → Or4 A B C D
- | Right3 : C → Or4 A B C D
- | Extra3: D → Or4 A B C D.
-
-interpretation "constructive ternary or" 'or4 x y z t = (Or4 x y z t).
-
-notation < "hvbox(a break ∨ b break ∨ c break ∨ d)" with precedence 35 for @{'or4 $a $b $c $d}.
-
-inductive And (A,B:CProp0) : CProp0 ≝
- | Conj : A → B → And A B.
-
-interpretation "constructive and" 'and x y = (And x y).
-
-inductive And3 (A,B,C:CProp0) : CProp0 ≝
- | Conj3 : A → B → C → And3 A B C.
-
-notation < "hvbox(a break ∧ b break ∧ c)" with precedence 35 for @{'and3 $a $b $c}.
-
-interpretation "constructive ternary and" 'and3 x y z = (And3 x y z).
-
-inductive And42 (A,B,C,D:CProp2) : CProp2 ≝
- | Conj42 : A → B → C → D → And42 A B C D.
-
-notation < "hvbox(a break ∧ b break ∧ c break ∧ d)" with precedence 35 for @{'and4 $a $b $c $d}.
-
-interpretation "constructive quaternary and2" 'and4 x y z t = (And42 x y z t).
-
-record Iff (A,B:CProp0) : CProp0 ≝
- { if: A → B;
- fi: B → A
- }.
-
-record Iff1 (A,B:CProp1) : CProp1 ≝
- { if1: A → B;
- fi1: B → A
- }.
-
-notation "hvbox(a break ⇔ b)" right associative with precedence 25 for @{'iff1 $a $b}.
-interpretation "logical iff" 'iff x y = (Iff x y).
-interpretation "logical iff type1" 'iff1 x y = (Iff1 x y).
-
-inductive exT22 (A:Type2) (P:A→CProp2) : CProp2 ≝
- ex_introT22: ∀w:A. P w → exT22 A P.
-
-interpretation "CProp2 exists" 'exists \eta.x = (exT22 ? x).
-
-definition pi1exT22 ≝ λA,P.λx:exT22 A P.match x with [ex_introT22 x _ ⇒ x].
-definition pi2exT22 ≝
- λA,P.λx:exT22 A P.match x return λx.P (pi1exT22 ?? x) with [ex_introT22 _ p ⇒ p].
-
-interpretation "exT22 \fst" 'pi1 = (pi1exT22 ? ?).
-interpretation "exT22 \snd" 'pi2 = (pi2exT22 ? ?).
-interpretation "exT22 \fst a" 'pi1a x = (pi1exT22 ? ? x).
-interpretation "exT22 \snd a" 'pi2a x = (pi2exT22 ? ? x).
-interpretation "exT22 \fst b" 'pi1b x y = (pi1exT22 ? ? x y).
-interpretation "exT22 \snd b" 'pi2b x y = (pi2exT22 ? ? x y).
-
-inductive exT (A:Type0) (P:A→CProp0) : CProp0 ≝
- ex_introT: ∀w:A. P w → exT A P.
-
-interpretation "CProp exists" 'exists \eta.x = (exT ? x).
-
-notation "\ll term 19 a, break term 19 b \gg"
-with precedence 90 for @{'dependent_pair $a $b}.
-interpretation "dependent pair" 'dependent_pair a b =
- (ex_introT ? ? a b).
-
-
-definition pi1exT ≝ λA,P.λx:exT A P.match x with [ex_introT x _ ⇒ x].
-definition pi2exT ≝
- λA,P.λx:exT A P.match x return λx.P (pi1exT ?? x) with [ex_introT _ p ⇒ p].
-
-interpretation "exT \fst" 'pi1 = (pi1exT ? ?).
-interpretation "exT \fst a" 'pi1a x = (pi1exT ? ? x).
-interpretation "exT \fst b" 'pi1b x y = (pi1exT ? ? x y).
-interpretation "exT \snd" 'pi2 = (pi2exT ? ?).
-interpretation "exT \snd a" 'pi2a x = (pi2exT ? ? x).
-interpretation "exT \snd b" 'pi2b x y = (pi2exT ? ? x y).
-
-inductive exT23 (A:Type0) (P:A→CProp0) (Q:A→CProp0) (R:A→A→CProp0) : CProp0 ≝
- ex_introT23: ∀w,p:A. P w → Q p → R w p → exT23 A P Q R.
-
-definition pi1exT23 ≝
- λA,P,Q,R.λx:exT23 A P Q R.match x with [ex_introT23 x _ _ _ _ ⇒ x].
-definition pi2exT23 ≝
- λA,P,Q,R.λx:exT23 A P Q R.match x with [ex_introT23 _ x _ _ _ ⇒ x].
-
-interpretation "exT2 \fst" 'pi1 = (pi1exT23 ? ? ? ?).
-interpretation "exT2 \snd" 'pi2 = (pi2exT23 ? ? ? ?).
-interpretation "exT2 \fst a" 'pi1a x = (pi1exT23 ? ? ? ? x).
-interpretation "exT2 \snd a" 'pi2a x = (pi2exT23 ? ? ? ? x).
-interpretation "exT2 \fst b" 'pi1b x y = (pi1exT23 ? ? ? ? x y).
-interpretation "exT2 \snd b" 'pi2b x y = (pi2exT23 ? ? ? ? x y).
-
-inductive exT2 (A:Type0) (P,Q:A→CProp0) : CProp0 ≝
- ex_introT2: ∀w:A. P w → Q w → exT2 A P Q.
-
-
-definition Not : CProp0 → Prop ≝ λx:CProp.x → False.
-
-interpretation "constructive not" 'not x = (Not x).
-
-definition cotransitive: ∀C:Type0. ∀lt:C→C→CProp0.CProp0 ≝
- λC:Type0.λlt:C→C→CProp0.∀x,y,z:C. lt x y → lt x z ∨ lt z y.
-
-definition coreflexive: ∀C:Type0. ∀lt:C→C→CProp0.CProp0 ≝
- λC:Type0.λlt:C→C→CProp0. ∀x:C. ¬ (lt x x).
-
-definition symmetric: ∀C:Type0. ∀lt:C→C→CProp0.CProp0 ≝
- λC:Type0.λlt:C→C→CProp0. ∀x,y:C.lt x y → lt y x.
-
-definition antisymmetric: ∀A:Type0. ∀R:A→A→CProp0. ∀eq:A→A→Prop.CProp0 ≝
- λA:Type0.λR:A→A→CProp0.λeq:A→A→Prop.∀x:A.∀y:A.R x y→R y x→eq x y.
-
-definition reflexive: ∀C:Type0. ∀lt:C→C→CProp0.CProp0 ≝ λA:Type0.λR:A→A→CProp0.∀x:A.R x x.
-
-definition transitive: ∀C:Type0. ∀lt:C→C→CProp0.CProp0 ≝ λA:Type0.λR:A→A→CProp0.∀x,y,z:A.R x y → R y z → R x z.
-
-definition reflexive1: ∀A:Type1.∀R:A→A→CProp1.CProp1 ≝ λA:Type1.λR:A→A→CProp1.∀x:A.R x x.
-definition symmetric1: ∀A:Type1.∀R:A→A→CProp1.CProp1 ≝ λC:Type1.λlt:C→C→CProp1. ∀x,y:C.lt x y → lt y x.
-definition transitive1: ∀A:Type1.∀R:A→A→CProp1.CProp1 ≝ λA:Type1.λR:A→A→CProp1.∀x,y,z:A.R x y → R y z → R x z.
-
-definition reflexive2: ∀A:Type2.∀R:A→A→CProp2.CProp2 ≝ λA:Type2.λR:A→A→CProp2.∀x:A.R x x.
-definition symmetric2: ∀A:Type2.∀R:A→A→CProp2.CProp2 ≝ λC:Type2.λlt:C→C→CProp2. ∀x,y:C.lt x y → lt y x.
-definition transitive2: ∀A:Type2.∀R:A→A→CProp2.CProp2 ≝ λA:Type2.λR:A→A→CProp2.∀x,y,z:A.R x y → R y z → R x z.
-
-definition reflexive3: ∀A:Type3.∀R:A→A→CProp3.CProp3 ≝ λA:Type3.λR:A→A→CProp3.∀x:A.R x x.
-definition symmetric3: ∀A:Type3.∀R:A→A→CProp3.CProp3 ≝ λC:Type3.λlt:C→C→CProp3. ∀x,y:C.lt x y → lt y x.
-definition transitive3: ∀A:Type3.∀R:A→A→CProp3.CProp3 ≝ λA:Type3.λR:A→A→CProp3.∀x,y,z:A.R x y → R y z → R x z.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/basic_topologies.ma".
-
-(*
-definition downarrow: ∀S:BTop. unary_morphism (Ω \sup S) (Ω \sup S).
- intros; constructor 1;
- [ apply (λU:Ω \sup S.{a | ∃b:carrbt S. b ∈ U ∧ a ∈ A ? (singleton ? b)});
- intros; simplify; split; intro; cases H1; cases x; exists [1,3: apply w]
- split; try assumption; [ apply (. H‡#) | apply (. H \sup -1‡#) ] assumption
- | intros; split; intros 2; cases f; exists; [1,3: apply w] cases x; split;
- try assumption; [ apply (. #‡H) | apply (. #‡H \sup -1)] assumption]
-qed.
-
-interpretation "downarrow" 'downarrow a = (fun_1 ?? (downarrow ?) a).
-
-definition ffintersects: ∀S:BTop. binary_morphism1 (Ω \sup S) (Ω \sup S) (Ω \sup S).
- intros; constructor 1;
- [ apply (λU,V. ↓U ∩ ↓V);
- | intros; apply (.= (†H)‡(†H1)); apply refl1]
-qed.
-
-interpretation "ffintersects" 'fintersects U V = (fun1 ??? (ffintersects ?) U V).
-
-record formal_topology: Type ≝
- { bt:> BTop;
- converges: ∀U,V: Ω \sup bt. A ? (U ↓ V) = A ? U ∩ A ? V
- }.
-
-
-definition ffintersects': ∀S:BTop. binary_morphism1 S S (Ω \sup S).
- intros; constructor 1;
- [ apply (λb,c:S. (singleton S b) ↓ (singleton S c));
- | intros; apply (.= (†H)‡(†H1)); apply refl1]
-qed.
-
-interpretation "ffintersects'" 'fintersects U V = (fun1 ??? (ffintersects' ?) U V).
-
-record formal_map (S,T: formal_topology) : Type ≝
- { cr:> continuous_relation_setoid S T;
- C1: ∀b,c. extS ?? cr (b ↓ c) = ext ?? cr b ↓ ext ?? cr c;
- C2: extS ?? cr T = S
- }.
-
-definition formal_map_setoid: formal_topology → formal_topology → setoid1.
- intros (S T); constructor 1;
- [ apply (formal_map S T);
- | constructor 1;
- [ apply (λf,f1: formal_map S T.f=f1);
- | simplify; intros 1; apply refl1
- | simplify; intros 2; apply sym1
- | simplify; intros 3; apply trans1]]
-qed.
-
-axiom C1':
- ∀S,T: formal_topology.∀f:formal_map_setoid S T.∀U,V: Ω \sup T.
- extS ?? f (U ↓ V) = extS ?? f U ↓ extS ?? f V.
-
-definition formal_map_composition:
- ∀o1,o2,o3: formal_topology.
- binary_morphism1
- (formal_map_setoid o1 o2)
- (formal_map_setoid o2 o3)
- (formal_map_setoid o1 o3).
- intros; constructor 1;
- [ intros; whd in c c1; constructor 1;
- [ apply (comp1 BTop ??? c c1);
- | intros;
- apply (.= (extS_com ??? c c1 ?));
- apply (.= †(C1 ?????));
- apply (.= (C1' ?????));
- apply (.= ((†((extS_singleton ????) \sup -1))‡(†((extS_singleton ????) \sup -1))));
- apply (.= (extS_com ??????)\sup -1 ‡ (extS_com ??????) \sup -1);
- apply (.= (extS_singleton ????)‡(extS_singleton ????));
- apply refl1;
- | apply (.= (extS_com ??? c c1 ?));
- apply (.= (†(C2 ???)));
- apply (.= (C2 ???));
- apply refl1;]
- | intros; simplify;
- change with (comp1 BTop ??? a b = comp1 BTop ??? a' b');
- apply prop1; assumption]
-qed.
-
-*)
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-notation "hvbox (r \sub \c)" with precedence 90 for @{'concr_rel $r}.
-notation "hvbox (r \sub \f)" with precedence 90 for @{'form_rel $r}.
-
-definition hide ≝ λA:Type.λx:A.x.
-
-interpretation "hide" 'hide x = (hide ? x).
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/categories.ma".
-
-inductive bool : Type0 := true : bool | false : bool.
-
-lemma BOOL : objs1 SET.
-constructor 1; [apply bool] constructor 1;
-[ intros (x y); apply (match x with [ true ⇒ match y with [ true ⇒ True | _ ⇒ False] | false ⇒ match y with [ true ⇒ False | false ⇒ True ]]);
-| whd; simplify; intros; cases x; apply I;
-| whd; simplify; intros 2; cases x; cases y; simplify; intros; assumption;
-| whd; simplify; intros 3; cases x; cases y; cases z; simplify; intros;
- try assumption; apply I]
-qed.
-
-lemma IF_THEN_ELSE_p :
- ∀S:setoid1.∀a,b:S.∀x,y:BOOL.x = y →
- (λm.match m with [ true ⇒ a | false ⇒ b ]) x =
- (λm.match m with [ true ⇒ a | false ⇒ b ]) y.
-whd in ⊢ (?→?→?→%→?);
-intros; cases x in e; cases y; simplify; intros; try apply refl1; whd in e; cases e;
-qed.
-
-interpretation "unary morphism comprehension with no proof" 'comprehension T P =
- (mk_unary_morphism T ? P ?).
-interpretation "unary morphism1 comprehension with no proof" 'comprehension T P =
- (mk_unary_morphism1 T ? P ?).
-
-notation > "hvbox({ ident i ∈ s | term 19 p | by })" with precedence 90
-for @{ 'comprehension_by $s (λ${ident i}. $p) $by}.
-notation < "hvbox({ ident i ∈ s | term 19 p })" with precedence 90
-for @{ 'comprehension_by $s (λ${ident i}:$_. $p) $by}.
-
-interpretation "unary morphism comprehension with proof" 'comprehension_by s \eta.f p =
- (mk_unary_morphism s ? f p).
-interpretation "unary morphism1 comprehension with proof" 'comprehension_by s \eta.f p =
- (mk_unary_morphism1 s ? f p).
-
-(* per il set-indexing vedere capitolo BPTools (foundational tools), Sect. 0.3.4 complete
- lattices, Definizione 0.9 *)
-(* USARE L'ESISTENZIALE DEBOLE *)
-
-definition if_then_else ≝ λT:Type.λe,t,f.match e return λ_.T with [ true ⇒ t | false ⇒ f].
-notation > "'If' term 19 e 'then' term 19 t 'else' term 90 f" non associative with precedence 19 for @{ 'if_then_else $e $t $f }.
-notation < "'If' \nbsp term 19 e \nbsp 'then' \nbsp term 19 t \nbsp 'else' \nbsp term 90 f \nbsp" non associative with precedence 19 for @{ 'if_then_else $e $t $f }.
-interpretation "Formula if_then_else" 'if_then_else e t f = (if_then_else ? e t f).
-
-notation > "hvbox(a break ≤ b)" non associative with precedence 45 for @{oa_leq $a $b}.
-notation > "a >< b" non associative with precedence 45 for @{oa_overlap $a $b}.
-notation > "⋁ p" non associative with precedence 45 for @{oa_join ? $p}.
-notation > "⋀ p" non associative with precedence 45 for @{oa_meet ? $p}.
-notation > "𝟙" non associative with precedence 90 for @{oa_one}.
-notation > "𝟘" non associative with precedence 90 for @{oa_zero}.
-record OAlgebra : Type2 := {
- oa_P :> SET1;
- oa_leq : oa_P × oa_P ⇒_1 CPROP;
- oa_overlap: oa_P × oa_P ⇒_1 CPROP;
- oa_meet: ∀I:SET.(I ⇒_2 oa_P) ⇒_2. oa_P;
- oa_join: ∀I:SET.(I ⇒_2 oa_P) ⇒_2. oa_P;
- oa_one: oa_P;
- oa_zero: oa_P;
- oa_leq_refl: ∀a:oa_P. a ≤ a;
- oa_leq_antisym: ∀a,b:oa_P.a ≤ b → b ≤ a → a = b;
- oa_leq_trans: ∀a,b,c:oa_P.a ≤ b → b ≤ c → a ≤ c;
- oa_overlap_sym: ∀a,b:oa_P.a >< b → b >< a;
- oa_meet_inf: ∀I:SET.∀p_i:I ⇒_2 oa_P.∀p:oa_P.p ≤ (⋀ p_i) = (∀i:I.p ≤ (p_i i));
- oa_join_sup: ∀I:SET.∀p_i:I ⇒_2 oa_P.∀p:oa_P.(⋁ p_i) ≤ p = (∀i:I.p_i i ≤ p);
- 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 | 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
- 2) manca (vedere altro capitolo) la "suriettivita'" come immagine di insiemi di oa_base
- oa_enum : ums oa_base oa_P;
- oa_density: ∀p,q.(∀i.oa_overlap p (oa_enum i) → oa_overlap q (oa_enum i)) → oa_leq p q
- *)
- oa_density: ∀p,q.(∀r.p >< r → q >< r) → p ≤ q
-}.
-
-notation "hvbox(a break ≤ b)" non associative with precedence 45 for @{ 'leq $a $b }.
-
-interpretation "o-algebra leq" 'leq a b = (fun21 ??? (oa_leq ?) a b).
-
-notation "hovbox(a mpadded width -150% (>)< b)" non associative with precedence 45
-for @{ 'overlap $a $b}.
-interpretation "o-algebra overlap" 'overlap a b = (fun21 ??? (oa_overlap ?) a b).
-
-notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∧) \below (\emsp) \nbsp term 90 p)"
-non associative with precedence 50 for @{ 'oa_meet $p }.
-notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∧) \below (ident i ∈ I) break term 90 p)"
-non associative with precedence 50 for @{ 'oa_meet_mk (λ${ident i}:$I.$p) }.
-
-notation > "hovbox(∧ f)" non associative with precedence 60
-for @{ 'oa_meet $f }.
-interpretation "o-algebra meet" 'oa_meet f =
- (fun12 ?? (oa_meet ??) f).
-interpretation "o-algebra meet with explicit function" 'oa_meet_mk f =
- (fun12 ?? (oa_meet ??) (mk_unary_morphism1 ?? f ?)).
-
-notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∨) \below (\emsp) \nbsp term 90 p)"
-non associative with precedence 50 for @{ 'oa_join $p }.
-notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∨) \below (ident i ∈ I) break term 90 p)"
-non associative with precedence 50 for @{ 'oa_join_mk (λ${ident i}:$I.$p) }.
-
-notation > "hovbox(∨ f)" non associative with precedence 60
-for @{ 'oa_join $f }.
-interpretation "o-algebra join" 'oa_join f =
- (fun12 ?? (oa_join ??) f).
-interpretation "o-algebra join with explicit function" 'oa_join_mk f =
- (fun12 ?? (oa_join ??) (mk_unary_morphism ?? f ?)).
-
-definition binary_meet : ∀O:OAlgebra. O × O ⇒_1 O.
-intros; split;
-[ intros (p q);
- apply (∧ { x ∈ BOOL | match x with [ true ⇒ p | false ⇒ q ] | IF_THEN_ELSE_p ? p q });
-| intros; lapply (prop12 ? O (oa_meet O BOOL));
- [2: apply ({ x ∈ BOOL | match x with [ true ⇒ a | false ⇒ b ] | IF_THEN_ELSE_p ? a b });
- |3: apply ({ x ∈ BOOL | match x with [ true ⇒ a' | false ⇒ b' ] | IF_THEN_ELSE_p ? a' b' });
- | apply Hletin;]
- intro x; simplify; cases x; simplify; assumption;]
-qed.
-
-interpretation "o-algebra binary meet" 'and a b =
- (fun21 ??? (binary_meet ?) a b).
-
-prefer coercion Type1_OF_OAlgebra.
-
-definition binary_join : ∀O:OAlgebra. O × O ⇒_1 O.
-intros; split;
-[ intros (p q);
- apply (∨ { x ∈ BOOL | match x with [ true ⇒ p | false ⇒ q ] | IF_THEN_ELSE_p ? p q });
-| intros; lapply (prop12 ? O (oa_join O BOOL));
- [2: apply ({ x ∈ BOOL | match x with [ true ⇒ a | false ⇒ b ] | IF_THEN_ELSE_p ? a b });
- |3: apply ({ x ∈ BOOL | match x with [ true ⇒ a' | false ⇒ b' ] | IF_THEN_ELSE_p ? a' b' });
- | apply Hletin;]
- intro x; simplify; cases x; simplify; assumption;]
-qed.
-
-interpretation "o-algebra binary join" 'or a b =
- (fun21 ??? (binary_join ?) a b).
-
-lemma oa_overlap_preservers_meet: ∀O:OAlgebra.∀p,q:O.p >< q → p >< (p ∧ q).
-intros; lapply (oa_overlap_preserves_meet_ O p q f) as H; clear f;
-(** screenshot "screenoa". *)
-assumption;
-qed.
-
-notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∨) \below (\emsp) \nbsp term 90 p)"
-non associative with precedence 49 for @{ 'oa_join $p }.
-notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∨) \below (ident i ∈ I) break term 90 p)"
-non associative with precedence 49 for @{ 'oa_join_mk (λ${ident i}:$I.$p) }.
-notation < "hovbox(a ∨ b)" left associative with precedence 49
-for @{ 'oa_join_mk (λ${ident i}:$_.match $i with [ true ⇒ $a | false ⇒ $b ]) }.
-
-notation > "hovbox(∨ f)" non associative with precedence 59
-for @{ 'oa_join $f }.
-notation > "hovbox(a ∨ b)" left associative with precedence 49
-for @{ 'oa_join (mk_unary_morphism BOOL ? (λx__:bool.match x__ with [ true ⇒ $a | false ⇒ $b ]) (IF_THEN_ELSE_p ? $a $b)) }.
-
-interpretation "o-algebra join" 'oa_join f =
- (fun12 ?? (oa_join ??) f).
-interpretation "o-algebra join with explicit function" 'oa_join_mk f =
- (fun12 ?? (oa_join ??) (mk_unary_morphism ?? f ?)).
-
-record ORelation (P,Q : OAlgebra) : Type2 ≝ {
- or_f_ : P ⇒_2 Q;
- or_f_minus_star_ : P ⇒_2 Q;
- or_f_star_ : Q ⇒_2 P;
- or_f_minus_ : Q ⇒_2 P;
- or_prop1_ : ∀p,q. (or_f_ p ≤ q) = (p ≤ or_f_star_ q);
- or_prop2_ : ∀p,q. (or_f_minus_ p ≤ q) = (p ≤ or_f_minus_star_ q);
- or_prop3_ : ∀p,q. (or_f_ p >< q) = (p >< or_f_minus_ q)
-}.
-
-definition ORelation_setoid : OAlgebra → OAlgebra → setoid2.
-intros (P Q);
-constructor 1;
-[ apply (ORelation P Q);
-| constructor 1;
- (* tenere solo una uguaglianza e usare la proposizione 9.9 per
- le altre (unicita' degli aggiunti e del simmetrico) *)
- [ apply (λp,q. And42
- (or_f_minus_star_ ?? p = or_f_minus_star_ ?? q)
- (or_f_minus_ ?? p = or_f_minus_ ?? q)
- (or_f_ ?? p = or_f_ ?? q)
- (or_f_star_ ?? p = or_f_star_ ?? q));
- | whd; simplify; intros; repeat split; intros; apply refl2;
- | whd; simplify; intros; cases a; clear a; split;
- intro a; apply sym1; generalize in match a;assumption;
- | whd; simplify; intros; cases a; cases a1; clear a a1; split; intro a;
- [ apply (.= (e a)); apply e4;
- | apply (.= (e1 a)); apply e5;
- | apply (.= (e2 a)); apply e6;
- | apply (.= (e3 a)); apply e7;]]]
-qed.
-
-definition ORelation_of_ORelation_setoid :
- ∀P,Q.ORelation_setoid P Q → ORelation P Q ≝ λP,Q,x.x.
-coercion ORelation_of_ORelation_setoid.
-
-definition or_f_minus_star: ∀P,Q:OAlgebra.(ORelation_setoid P Q) ⇒_2 (P ⇒_2 Q).
- intros; constructor 1;
- [ apply or_f_minus_star_;
- | intros; cases e; assumption]
-qed.
-
-definition or_f: ∀P,Q:OAlgebra.(ORelation_setoid P Q) ⇒_2 (P ⇒_2 Q).
- intros; constructor 1;
- [ apply or_f_;
- | intros; cases e; assumption]
-qed.
-
-definition or_f_minus: ∀P,Q:OAlgebra.(ORelation_setoid P Q) ⇒_2 (Q ⇒_2 P).
- intros; constructor 1;
- [ apply or_f_minus_;
- | intros; cases e; assumption]
-qed.
-
-definition or_f_star: ∀P,Q:OAlgebra.(ORelation_setoid P Q) ⇒_2 (Q ⇒_2 P).
- intros; constructor 1;
- [ apply or_f_star_;
- | intros; cases e; assumption]
-qed.
-
-lemma arrows1_of_ORelation_setoid : ∀P,Q. ORelation_setoid P Q → (P ⇒_2 Q).
-intros; apply (or_f ?? c);
-qed.
-coercion arrows1_of_ORelation_setoid.
-
-interpretation "o-relation f⎻*" 'OR_f_minus_star r = (fun12 ?? (or_f_minus_star ? ?) r).
-interpretation "o-relation f⎻" 'OR_f_minus r = (fun12 ?? (or_f_minus ? ?) r).
-interpretation "o-relation f*" 'OR_f_star r = (fun12 ?? (or_f_star ? ?) r).
-
-definition or_prop1 : ∀P,Q:OAlgebra.∀F:ORelation_setoid P Q.∀p,q.
- (F p ≤ q) =_1 (p ≤ F* q).
-intros; apply (or_prop1_ ?? F p q);
-qed.
-
-definition or_prop2 : ∀P,Q:OAlgebra.∀F:ORelation_setoid P Q.∀p,q.
- (F⎻ p ≤ q) = (p ≤ F⎻* q).
-intros; apply (or_prop2_ ?? F p q);
-qed.
-
-definition or_prop3 : ∀P,Q:OAlgebra.∀F:ORelation_setoid P Q.∀p,q.
- (F p >< q) = (p >< F⎻ q).
-intros; apply (or_prop3_ ?? F p q);
-qed.
-
-definition ORelation_composition : ∀P,Q,R.
- (ORelation_setoid P Q) × (ORelation_setoid Q R) ⇒_2 (ORelation_setoid P R).
-intros;
-constructor 1;
-[ intros (F G);
- constructor 1;
- [ apply (G ∘ F);
- | apply rule (G⎻* ∘ F⎻* );
- | apply (F* ∘ G* );
- | apply (F⎻ ∘ G⎻);
- | intros;
- change with ((G (F p) ≤ q) = (p ≤ (F* (G* q))));
- apply (.= (or_prop1 :?));
- apply (or_prop1 :?);
- | intros;
- change with ((F⎻ (G⎻ p) ≤ q) = (p ≤ (G⎻* (F⎻* q))));
- apply (.= (or_prop2 :?));
- apply or_prop2 ;
- | intros; change with ((G (F (p)) >< q) = (p >< (F⎻ (G⎻ q))));
- apply (.= (or_prop3 :?));
- apply or_prop3;
- ]
-| intros; split; simplify;
- [3: unfold arrows1_of_ORelation_setoid; apply ((†e)‡(†e1));
- |1: apply ((†e)‡(†e1));
- |2,4: apply ((†e1)‡(†e));]]
-qed.
-
-definition OA : category2.
-split;
-[ apply (OAlgebra);
-| intros; apply (ORelation_setoid o o1);
-| intro O; split;
- [1,2,3,4: apply id2;
- |5,6,7:intros; apply refl1;]
-| apply ORelation_composition;
-| intros (P Q R S F G H); split;
- [ change with (H⎻* ∘ G⎻* ∘ F⎻* = H⎻* ∘ (G⎻* ∘ F⎻* ));
- apply (comp_assoc2 ????? (F⎻* ) (G⎻* ) (H⎻* ));
- | apply ((comp_assoc2 ????? (H⎻) (G⎻) (F⎻))^-1);
- | apply ((comp_assoc2 ????? F G H)^-1);
- | apply ((comp_assoc2 ????? H* G* F* ));]
-| intros; split; unfold ORelation_composition; simplify; apply id_neutral_left2;
-| intros; split; unfold ORelation_composition; simplify; apply id_neutral_right2;]
-qed.
-
-definition OAlgebra_of_objs2_OA: objs2 OA → OAlgebra ≝ λx.x.
-coercion OAlgebra_of_objs2_OA.
-
-definition ORelation_setoid_of_arrows2_OA:
- ∀P,Q. arrows2 OA P Q → ORelation_setoid P Q ≝ λP,Q,c.c.
-coercion ORelation_setoid_of_arrows2_OA.
-
-prefer coercion Type_OF_objs2.
-
-notation > "B ⇒_\o2 C" right associative with precedence 72 for @{'arrows2_OA $B $C}.
-notation "B ⇒\sub (\o 2) C" right associative with precedence 72 for @{'arrows2_OA $B $C}.
-interpretation "'arrows2_OA" 'arrows2_OA A B = (arrows2 OA A B).
-
-(* qui la notazione non va *)
-lemma leq_to_eq_join: ∀S:OA.∀p,q:S. p ≤ q → q = (binary_join ? p q).
- intros;
- apply oa_leq_antisym;
- [ apply oa_density; intros;
- apply oa_overlap_sym;
- unfold binary_join; simplify;
- apply (. (oa_join_split : ?));
- exists; [ apply false ]
- apply oa_overlap_sym;
- assumption
- | unfold binary_join; simplify;
- apply (. (oa_join_sup : ?)); intro;
- cases i; whd in ⊢ (? ? ? ? ? % ?);
- [ assumption | apply oa_leq_refl ]]
-qed.
-
-lemma overlap_monotone_left: ∀S:OA.∀p,q,r:S. p ≤ q → p >< r → q >< r.
- intros;
- apply (. (leq_to_eq_join : ?)‡#);
- [ apply f;
- | skip
- | apply oa_overlap_sym;
- unfold binary_join; simplify;
- apply (. (oa_join_split : ?));
- exists [ apply true ]
- apply oa_overlap_sym;
- assumption; ]
-qed.
-
-(* Part of proposition 9.9 *)
-lemma f_minus_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R⎻ p ≤ R⎻ q.
- intros;
- apply (. (or_prop2 : ?));
- apply oa_leq_trans; [2: apply f; | skip | apply (. (or_prop2 : ?)^ -1); apply oa_leq_refl;]
-qed.
-
-(* Part of proposition 9.9 *)
-lemma f_minus_star_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R⎻* p ≤ R⎻* q.
- intros;
- apply (. (or_prop2 : ?)^ -1);
- apply oa_leq_trans; [3: apply f; | skip | apply (. (or_prop2 : ?)); apply oa_leq_refl;]
-qed.
-
-(* Part of proposition 9.9 *)
-lemma f_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R p ≤ R q.
- intros;
- apply (. (or_prop1 : ?));
- apply oa_leq_trans; [2: apply f; | skip | apply (. (or_prop1 : ?)^ -1); apply oa_leq_refl;]
-qed.
-
-(* Part of proposition 9.9 *)
-lemma f_star_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R* p ≤ R* q.
- intros;
- apply (. (or_prop1 : ?)^ -1);
- apply oa_leq_trans; [3: apply f; | skip | apply (. (or_prop1 : ?)); apply oa_leq_refl;]
-qed.
-
-lemma lemma_10_2_a: ∀S,T.∀R:arrows2 OA S T.∀p. p ≤ R⎻* (R⎻ p).
- intros;
- apply (. (or_prop2 : ?)^-1);
- apply oa_leq_refl.
-qed.
-
-lemma lemma_10_2_b: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻ (R⎻* p) ≤ p.
- intros;
- apply (. (or_prop2 : ?));
- apply oa_leq_refl.
-qed.
-
-lemma lemma_10_2_c: ∀S,T.∀R:arrows2 OA S T.∀p. p ≤ R* (R p).
- intros;
- apply (. (or_prop1 : ?)^-1);
- apply oa_leq_refl.
-qed.
-
-lemma lemma_10_2_d: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* p) ≤ p.
- intros;
- apply (. (or_prop1 : ?));
- apply oa_leq_refl.
-qed.
-
-lemma lemma_10_3_a: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻ (R⎻* (R⎻ p)) = R⎻ p.
- intros; apply oa_leq_antisym;
- [ apply lemma_10_2_b;
- | apply f_minus_image_monotone;
- apply lemma_10_2_a; ]
-qed.
-
-lemma lemma_10_3_b: ∀S,T.∀R:arrows2 OA S T.∀p. R* (R (R* p)) = R* p.
- intros; apply oa_leq_antisym;
- [ apply f_star_image_monotone;
- apply (lemma_10_2_d ?? R p);
- | apply lemma_10_2_c; ]
-qed.
-
-lemma lemma_10_3_c: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* (R p)) = R p.
- intros; apply oa_leq_antisym;
- [ apply lemma_10_2_d;
- | apply f_image_monotone;
- apply (lemma_10_2_c ?? R p); ]
-qed.
-
-lemma lemma_10_3_d: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻* (R⎻ (R⎻* p)) = R⎻* p.
- intros; apply oa_leq_antisym;
- [ apply f_minus_star_image_monotone;
- apply (lemma_10_2_b ?? R p);
- | apply lemma_10_2_a; ]
-qed.
-
-lemma lemma_10_4_a: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻* (R⎻ (R⎻* (R⎻ p))) = R⎻* (R⎻ p).
- intros; apply (†(lemma_10_3_a ?? R p));
-qed.
-
-lemma lemma_10_4_b: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* (R (R* p))) = R (R* p).
-intros; unfold in ⊢ (? ? ? % %); apply (†(lemma_10_3_b ?? R p));
-qed.
-
-lemma oa_overlap_sym': ∀o:OA.∀U,V:o. (U >< V) = (V >< U).
- intros; split; intro; apply oa_overlap_sym; assumption.
-qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/o-algebra.ma".
-include "formal_topology/notation.ma".
-
-record Obasic_pair: Type2 ≝ {
- Oconcr: OA; Oform: OA; Orel: arrows2 ? Oconcr Oform
-}.
-
-(* FIX *)
-interpretation "o-basic pair relation indexed" 'Vdash2 x y c = (Orel c x y).
-interpretation "o-basic pair relation (non applied)" 'Vdash c = (Orel c).
-
-record Orelation_pair (BP1,BP2: Obasic_pair): Type2 ≝ {
- Oconcr_rel: (Oconcr BP1) ⇒_\o2 (Oconcr BP2); Oform_rel: (Oform BP1) ⇒_\o2 (Oform BP2);
- Ocommute: ⊩ ∘ Oconcr_rel =_2 Oform_rel ∘ ⊩
-}.
-
-(* FIX *)
-interpretation "o-concrete relation" 'concr_rel r = (Oconcr_rel ?? r).
-interpretation "o-formal relation" 'form_rel r = (Oform_rel ?? r).
-
-definition Orelation_pair_equality:
- ∀o1,o2. equivalence_relation2 (Orelation_pair o1 o2).
- intros;
- constructor 1;
- [ apply (λr,r'. ⊩ ∘ r \sub\c = ⊩ ∘ r' \sub\c);
- | simplify;
- intros;
- apply refl2;
- | simplify;
- intros 2;
- apply sym2;
- | simplify;
- intros 3;
- apply trans2;
- ]
-qed.
-
-(* qui setoid1 e' giusto: ma non lo e'!!! *)
-definition Orelation_pair_setoid: Obasic_pair → Obasic_pair → setoid2.
- intros;
- constructor 1;
- [ apply (Orelation_pair o o1)
- | apply Orelation_pair_equality
- ]
-qed.
-
-definition Orelation_pair_of_Orelation_pair_setoid:
- ∀P,Q. Orelation_pair_setoid P Q → Orelation_pair P Q ≝ λP,Q,x.x.
-coercion Orelation_pair_of_Orelation_pair_setoid.
-
-lemma eq_to_eq': ∀o1,o2.∀r,r': Orelation_pair_setoid o1 o2. r =_2 r' → r \sub\f ∘ ⊩ =_2 r'\sub\f ∘ ⊩.
- intros 5 (o1 o2 r r' H); change in H with (⊩ ∘ r\sub\c = ⊩ ∘ r'\sub\c);
- apply (.= ((Ocommute ?? r) ^ -1));
- apply (.= H);
- apply (.= (Ocommute ?? r'));
- apply refl2;
-qed.
-
-
-definition Oid_relation_pair: ∀o:Obasic_pair. Orelation_pair o o.
- intro;
- constructor 1;
- [1,2: apply id2;
- | lapply (id_neutral_right2 ? (Oconcr o) ? (⊩)) as H;
- lapply (id_neutral_left2 ?? (Oform o) (⊩)) as H1;
- apply (.= H);
- apply (H1^-1);]
-qed.
-
-lemma Orelation_pair_composition:
- ∀o1,o2,o3:Obasic_pair.
- Orelation_pair_setoid o1 o2 → Orelation_pair_setoid o2 o3→Orelation_pair_setoid o1 o3.
-intros 3 (o1 o2 o3);
- intros (r r1);
- constructor 1;
- [ apply (r1 \sub\c ∘ r \sub\c)
- | apply (r1 \sub\f ∘ r \sub\f)
- | lapply (Ocommute ?? r) as H;
- lapply (Ocommute ?? r1) as H1;
- apply rule (.= ASSOC);
- apply (.= #‡H1);
- apply rule (.= ASSOC ^ -1);
- apply (.= H‡#);
- apply rule ASSOC]
-qed.
-
-
-lemma Orelation_pair_composition_is_morphism:
- ∀o1,o2,o3:Obasic_pair.
- Πa,a':Orelation_pair_setoid o1 o2.Πb,b':Orelation_pair_setoid o2 o3.
- a=a' →b=b' →
- Orelation_pair_composition o1 o2 o3 a b
- = Orelation_pair_composition o1 o2 o3 a' b'.
-intros;
- change with (⊩ ∘ (b\sub\c ∘ a\sub\c) = ⊩ ∘ (b'\sub\c ∘ a'\sub\c));
- change in e with (⊩ ∘ a \sub\c = ⊩ ∘ a' \sub\c);
- change in e1 with (⊩ ∘ b \sub\c = ⊩ ∘ b' \sub\c);
- apply rule (.= ASSOC);
- apply (.= #‡e1);
- apply (.= #‡(Ocommute ?? b'));
- apply rule (.= ASSOC^-1);
- apply (.= e‡#);
- apply rule (.= ASSOC);
- apply (.= #‡(Ocommute ?? b')^-1);
- apply rule (ASSOC^-1);
-qed.
-
-definition Orelation_pair_composition_morphism:
- ∀o1,o2,o3. binary_morphism2 (Orelation_pair_setoid o1 o2) (Orelation_pair_setoid o2 o3) (Orelation_pair_setoid o1 o3).
-intros; constructor 1;
-[ apply Orelation_pair_composition;
-| apply Orelation_pair_composition_is_morphism;]
-qed.
-
-lemma Orelation_pair_composition_morphism_assoc:
-∀o1,o2,o3,o4:Obasic_pair
- .Πa12:Orelation_pair_setoid o1 o2
- .Πa23:Orelation_pair_setoid o2 o3
- .Πa34:Orelation_pair_setoid o3 o4
- .Orelation_pair_composition_morphism o1 o3 o4
- (Orelation_pair_composition_morphism o1 o2 o3 a12 a23) a34
- =Orelation_pair_composition_morphism o1 o2 o4 a12
- (Orelation_pair_composition_morphism o2 o3 o4 a23 a34).
- intros;
- change with (⊩ ∘ (a34\sub\c ∘ (a23\sub\c ∘ a12\sub\c)) =
- ⊩ ∘ ((a34\sub\c ∘ a23\sub\c) ∘ a12\sub\c));
- apply rule (ASSOC‡#);
-qed.
-
-lemma Orelation_pair_composition_morphism_respects_id:
-Πo1:Obasic_pair
-.Πo2:Obasic_pair
- .Πa:Orelation_pair_setoid o1 o2
- .Orelation_pair_composition_morphism o1 o1 o2 (Oid_relation_pair o1) a=a.
- intros;
- change with (⊩ ∘ (a\sub\c ∘ (Oid_relation_pair o1)\sub\c) = ⊩ ∘ a\sub\c);
- apply ((id_neutral_right2 ????)‡#);
-qed.
-
-lemma Orelation_pair_composition_morphism_respects_id_r:
-Πo1:Obasic_pair
-.Πo2:Obasic_pair
- .Πa:Orelation_pair_setoid o1 o2
- .Orelation_pair_composition_morphism o1 o2 o2 a (Oid_relation_pair o2)=a.
-intros;
- change with (⊩ ∘ ((Oid_relation_pair o2)\sub\c ∘ a\sub\c) = ⊩ ∘ a\sub\c);
- apply ((id_neutral_left2 ????)‡#);
-qed.
-
-definition OBP: category2.
- constructor 1;
- [ apply Obasic_pair
- | apply Orelation_pair_setoid
- | apply Oid_relation_pair
- | apply Orelation_pair_composition_morphism
- | apply Orelation_pair_composition_morphism_assoc;
- | apply Orelation_pair_composition_morphism_respects_id;
- | apply Orelation_pair_composition_morphism_respects_id_r;]
-qed.
-
-definition Obasic_pair_of_objs2_OBP: objs2 OBP → Obasic_pair ≝ λx.x.
-coercion Obasic_pair_of_objs2_OBP.
-
-definition Orelation_pair_setoid_of_arrows2_OBP:
- ∀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;
- [ apply (ext ? ? (rel o));
- | intros;
- apply (.= #‡H);
- apply refl1]
-qed.
-
-definition BPextS: ∀o: BP. Ω \sup (form o) ⇒ Ω \sup (concr o) ≝
- λo.extS ?? (rel o).
-*)
-
-(*
-definition fintersects: ∀o: BP. binary_morphism1 (form o) (form o) (Ω \sup (form o)).
- intros (o); constructor 1;
- [ apply (λa,b: form o.{c | BPext o c ⊆ BPext o a ∩ BPext o b });
- intros; simplify; apply (.= (†H)‡#); apply refl1
- | intros; split; simplify; intros;
- [ apply (. #‡((†H)‡(†H1))); assumption
- | apply (. #‡((†H\sup -1)‡(†H1\sup -1))); assumption]]
-qed.
-
-interpretation "fintersects" 'fintersects U V = (fun1 ??? (fintersects ?) U V).
-
-definition fintersectsS:
- ∀o:BP. binary_morphism1 (Ω \sup (form o)) (Ω \sup (form o)) (Ω \sup (form o)).
- intros (o); constructor 1;
- [ apply (λo: basic_pair.λa,b: Ω \sup (form o).{c | BPext o c ⊆ BPextS o a ∩ BPextS o b });
- intros; simplify; apply (.= (†H)‡#); apply refl1
- | intros; split; simplify; intros;
- [ apply (. #‡((†H)‡(†H1))); assumption
- | apply (. #‡((†H\sup -1)‡(†H1\sup -1))); assumption]]
-qed.
-
-interpretation "fintersectsS" 'fintersects U V = (fun1 ??? (fintersectsS ?) U V).
-*)
-
-(*
-definition relS: ∀o: BP. binary_morphism1 (concr o) (Ω \sup (form o)) CPROP.
- intros (o); constructor 1;
- [ apply (λx:concr o.λS: Ω \sup (form o).∃y: form o.y ∈ S ∧ x ⊩ y);
- | intros; split; intros; cases H2; exists [1,3: apply w]
- [ apply (. (#‡H1)‡(H‡#)); assumption
- | apply (. (#‡H1 \sup -1)‡(H \sup -1‡#)); assumption]]
-qed.
-
-interpretation "basic pair relation for subsets" 'Vdash2 x y = (fun1 (concr ?) ?? (relS ?) x y).
-interpretation "basic pair relation for subsets (non applied)" 'Vdash = (fun1 ??? (relS ?)).
-*)
-
-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 = (fun12 ? ? (or_f_minus_star ? ?) (Orel 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 = (fun12 ? ? (or_f ? ?) (Orel 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 = (fun12 ? ? (or_f_star ? ?) (Orel 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 = (fun12 ? ? (or_f_minus ? ?) (Orel x)).
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/notation.ma".
-include "formal_topology/o-basic_pairs.ma".
-include "formal_topology/o-basic_topologies.ma".
-
-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 → OBTop.
- intro t;
- constructor 1;
- [ apply (Oform t);
- | apply (□⎽t ∘ Ext⎽t);
- | apply (◊⎽t ∘ Rest⎽t);
- | apply hide; intros 2; split; intro;
- [ change with ((⊩) \sup ⎻* ((⊩) \sup ⎻ U) ≤ (⊩) \sup ⎻* ((⊩) \sup ⎻ V));
- apply (. (#‡(lemma_10_4_a ?? (⊩) V)^-1));
- apply f_minus_star_image_monotone;
- apply f_minus_image_monotone;
- assumption
- | apply oa_leq_trans;
- [3: apply f;
- | skip
- | change with (U ≤ (⊩)⎻* ((⊩)⎻ U));
- apply (. (or_prop2 : ?) ^ -1);
- apply oa_leq_refl; ]]
- | apply hide; intros 2; split; intro;
- [ change with (◊⎽t ((⊩) \sup * U) ≤ ◊⎽t ((⊩) \sup * V));
- apply (. ((lemma_10_4_b ?? (⊩) U)^-1)‡#);
- apply (f_image_monotone ?? (⊩) ? ((⊩)* V));
- apply f_star_image_monotone;
- assumption;
- | apply oa_leq_trans;
- [2: apply f;
- | skip
- | change with ((⊩) ((⊩)* V) ≤ V);
- apply (. (or_prop1 : ?));
- apply oa_leq_refl; ]]
- | apply hide; intros;
- apply (.= (oa_overlap_sym' : ?));
- change with ((◊⎽t ((⊩)* V) >< (⊩)⎻* ((⊩)⎻ U)) = (U >< (◊⎽t ((⊩)* V))));
- apply (.= (or_prop3 ?? (⊩) ((⊩)* V) ?));
- apply (.= #‡(lemma_10_3_a : ?));
- apply (.= (or_prop3 : ?)^-1);
- apply (oa_overlap_sym' ? ((⊩) ((⊩)* V)) U); ]
-qed.
-
-definition o_continuous_relation_of_o_relation_pair:
- ∀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);
- | apply hide; unfold o_basic_topology_of_o_basic_pair; simplify; intros (U e);
- apply sym1;
- apply (.= †(†e));
- change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f ∘ (⊩)) ((⊩\sub BP1)* U));
- cut ((t \sub \f ∘ (⊩)) ((⊩\sub BP1)* U) = ((⊩) ∘ t \sub \c) ((⊩\sub BP1)* U)) as COM;[2:
- cases (Ocommute ?? t); apply (e3 ^ -1 ((⊩\sub BP1)* U));]
- apply (.= †COM);
- change in ⊢ (? ? ? % ?) with (((⊩) ∘ (⊩)* ) (((⊩) ∘ t \sub \c ∘ (⊩)* ) U));
- apply (.= (lemma_10_3_c ?? (⊩) (t \sub \c ((⊩\sub BP1)* U))));
- apply (.= COM ^ -1);
- change in ⊢ (? ? ? % ?) with (t \sub \f (((⊩) ∘ (⊩\sub BP1)* ) U));
- change in e with (U=((⊩)∘(⊩ \sub BP1) \sup * ) U);
- apply (†e^-1);
- | apply hide; unfold o_basic_topology_of_o_basic_pair; simplify; intros;
- apply sym1;
- apply (.= †(†e));
- change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f⎻* ∘ (⊩)⎻* ) ((⊩\sub BP1)⎻ U));
- cut ((t \sub \f⎻* ∘ (⊩)⎻* ) ((⊩\sub BP1)⎻ U) = ((⊩)⎻* ∘ t \sub \c⎻* ) ((⊩\sub BP1)⎻ U)) as COM;[2:
- cases (Ocommute ?? t); apply (e1 ^ -1 ((⊩\sub BP1)⎻ U));]
- apply (.= †COM);
- change in ⊢ (? ? ? % ?) with (((⊩)⎻* ∘ (⊩)⎻ ) (((⊩)⎻* ∘ t \sub \c⎻* ∘ (⊩)⎻ ) U));
- apply (.= (lemma_10_3_d ?? (⊩) (t \sub \c⎻* ((⊩\sub BP1)⎻ U))));
- apply (.= COM ^ -1);
- change in ⊢ (? ? ? % ?) with (t \sub \f⎻* (((⊩)⎻* ∘ (⊩\sub BP1)⎻ ) U));
- change in e with (U=((⊩)⎻* ∘(⊩ \sub BP1)⎻ ) U);
- apply (†e^-1);]
-qed.
-
-
-definition OR : carr3 (OBP ⇒_\c3 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 ⎻* ∘ 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 (.= #‡†(Ocommute:?)^-1);
- apply (.= #‡e1);
- change in ⊢ (? ? ? (? ? ? ? ? ? %) ?) with (⊩\sub T ∘ a'\sub\c)⎻*;
- apply (.= #‡†(Ocommute:?));
- change in ⊢ (? ? ? (? ? ? ? ? ? %) ?) with (a'\sub\f⎻* ∘ (⊩\sub S)⎻* );
- apply (.= (comp_assoc2 ? ???? ?? a'\sub\f⎻* )^-1);
- apply refl2;]
-| intros 2 (o a); apply refl1;
-| intros 6; apply refl1;]
-qed.
-
+++ /dev/null
- (**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/o-algebra.ma".
-include "formal_topology/o-saturations.ma".
-
-record Obasic_topology: Type2 ≝ {
- Ocarrbt:> OA;
- oA: Ocarrbt ⇒_2 Ocarrbt; oJ: Ocarrbt ⇒_2 Ocarrbt;
- oA_is_saturation: is_o_saturation ? oA; oJ_is_reduction: is_o_reduction ? oJ;
- Ocompatibility: ∀U,V. (oA U >< oJ V) =_1 (U >< oJ V)
- }.
-
-record Ocontinuous_relation (S,T: Obasic_topology) : Type2 ≝ {
- Ocont_rel:> arrows2 OA S T;
- Oreduced: ∀U:S. U = oJ ? U → Ocont_rel U =_1 oJ ? (Ocont_rel U);
- Osaturated: ∀U:S. U = oA ? U → Ocont_rel⎻* U =_1 oA ? (Ocont_rel⎻* U)
- }.
-
-definition Ocontinuous_relation_setoid: Obasic_topology → Obasic_topology → setoid2.
- intros (S T); constructor 1;
- [ apply (Ocontinuous_relation S T)
- | constructor 1;
- [ alias symbol "eq" = "setoid2 eq".
- alias symbol "compose" = "category2 composition".
- 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 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':
- ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
- a = a' → ∀X.a⎻* (A o1 X) = a'⎻* (A o1 X).
- intros; apply oa_leq_antisym; intro; unfold minus_star_image; simplify; intros;
- [ cut (ext ?? a a1 ⊆ A ? X); [2: intros 2; apply (H1 a2); cases f1; assumption;]
- lapply (if ?? (A_is_saturation ???) Hcut); clear Hcut;
- cut (A ? (ext ?? a' a1) ⊆ A ? X); [2: apply (. (H ?)‡#); assumption]
- lapply (fi ?? (A_is_saturation ???) Hcut);
- apply (Hletin1 x); change with (x ∈ ext ?? a' a1); split; simplify;
- [ apply I | assumption ]
- | cut (ext ?? a' a1 ⊆ A ? X); [2: intros 2; apply (H1 a2); cases f1; assumption;]
- lapply (if ?? (A_is_saturation ???) Hcut); clear Hcut;
- cut (A ? (ext ?? a a1) ⊆ A ? X); [2: apply (. (H ?)\sup -1‡#); assumption]
- lapply (fi ?? (A_is_saturation ???) Hcut);
- apply (Hletin1 x); change with (x ∈ ext ?? a a1); split; simplify;
- [ apply I | assumption ]]
-qed.
-
-theorem continuous_relation_eq_inv':
- ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
- (∀X.a⎻* (A o1 X) = a'⎻* (A o1 X)) → a=a'.
- intros 6;
- cut (∀a,a': continuous_relation_setoid o1 o2.
- (∀X.a⎻* (A o1 X) = a'⎻* (A o1 X)) →
- ∀V:(oa_P (carrbt o2)). A o1 (a'⎻ V) ≤ A o1 (a⎻ V));
- [2: clear b H a' a; intros;
- lapply depth=0 (λV.saturation_expansive ??? (extS ?? a V)); [2: apply A_is_saturation;|skip]
- (* fundamental adjunction here! to be taken out *)
- cut (∀V:Ω \sup o2.V ⊆ minus_star_image ?? a (A ? (extS ?? a V)));
- [2: intro; intros 2; unfold minus_star_image; simplify; intros;
- apply (Hletin V1 x); whd; split; [ exact I | exists; [apply a1] split; assumption]]
- clear Hletin;
- cut (∀V:Ω \sup o2.V ⊆ minus_star_image ?? a' (A ? (extS ?? a V)));
- [2: intro; apply (. #‡(H ?)); apply Hcut] clear H Hcut;
- (* second half of the fundamental adjunction here! to be taken out too *)
- intro; lapply (Hcut1 (singleton ? V)); clear Hcut1;
- unfold minus_star_image in Hletin; unfold singleton in Hletin; simplify in Hletin;
- whd in Hletin; whd in Hletin:(?→?→%); simplify in Hletin;
- apply (if ?? (A_is_saturation ???));
- intros 2 (x H); lapply (Hletin V ? x ?);
- [ apply refl | cases H; assumption; ]
- change with (x ∈ A ? (ext ?? a V));
- apply (. #‡(†(extS_singleton ????)));
- assumption;]
- split; apply Hcut; [2: assumption | intro; apply sym1; apply H]
-qed.
-*)
-
-definition Ocontinuous_relation_comp:
- ∀o1,o2,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 (.= (Oreduced : ?)^-1);
- [ apply (.= (Oreduced :?)); [ assumption | apply refl1 ]
- | apply refl1]
- | intros;
- apply sym1;
- change in match ((s ∘ r)⎻* U) with (s⎻* (r⎻* U));
- apply (.= (Osaturated : ?)^-1);
- [ apply (.= (Osaturated : ?)); [ assumption | apply refl1 ]
- | apply refl1]]
-qed.
-
-definition OBTop: category2.
- constructor 1;
- [ apply Obasic_topology
- | apply Ocontinuous_relation_setoid
- | intro; constructor 1;
- [ apply id2
- | intros; apply e;
- | intros; apply e;]
- | intros; constructor 1;
- [ apply Ocontinuous_relation_comp;
- | intros; simplify;
- 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 (b⎻* (a'⎻* (oA o1 x)) =_1 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 (b'⎻* (oA o2 (a'⎻* (oA o1 x))) =_1 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⎻* ) ∘ oA o1 = ((a34⎻* ∘ (a23⎻* ∘ a12⎻* )) ∘ oA o1));
- apply rule (#‡ASSOC ^ -1);
- | intros; simplify;
- change with ((a⎻* ∘ (id2 ? o1)⎻* ) ∘ oA o1 = a⎻* ∘ oA o1);
- apply (#‡(id_neutral_right2 : ?));
- | intros; simplify;
- change with (((id2 ? o2)⎻* ∘ a⎻* ) ∘ oA o1 = a⎻* ∘ oA o1);
- apply (#‡(id_neutral_left2 : ?));]
-qed.
-
-definition Obasic_topology_of_OBTop: objs2 OBTop → Obasic_topology ≝ λx.x.
-coercion Obasic_topology_of_OBTop.
-
-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.
-
-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 *)
-theorem continuous_relation_eqS:
- ∀o1,o2:basic_topology.∀a,a': continuous_relation_setoid o1 o2.
- a = a' → ∀X. A ? (extS ?? a X) = A ? (extS ?? a' X).
- intros;
- cut (∀a: arrows1 ? o1 ?.∀x. x ∈ extS ?? a X → ∃y:o2.y ∈ X ∧ x ∈ ext ?? a y);
- [2: intros; cases f; clear f; cases H1; exists [apply w] cases x1; split;
- try assumption; split; assumption]
- cut (∀a,a':continuous_relation_setoid o1 o2.eq1 ? a a' → ∀x. x ∈ extS ?? a X → ∃y:o2. y ∈ X ∧ x ∈ A ? (ext ?? a' y));
- [2: intros; cases (Hcut ?? f); exists; [apply w] cases x1; split; try assumption;
- apply (. #‡(H1 ?));
- apply (saturation_expansive ?? (A_is_saturation o1) (ext ?? a1 w) x);
- assumption;] clear Hcut;
- split; apply (if ?? (A_is_saturation ???)); intros 2;
- [lapply (Hcut1 a a' H a1 f) | lapply (Hcut1 a' a (H \sup -1) a1 f)]
- cases Hletin; clear Hletin; cases x; clear x;
- cut (∀a: arrows1 ? o1 ?. ext ?? a w ⊆ extS ?? a X);
- [2,4: intros 3; cases f3; clear f3; simplify in f5; split; try assumption;
- exists [1,3: apply w] split; assumption;]
- cut (∀a. A ? (ext o1 o2 a w) ⊆ A ? (extS o1 o2 a X));
- [2,4: intros; apply saturation_monotone; try (apply A_is_saturation); apply Hcut;]
- apply Hcut2; assumption.
-qed.
-*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/o-basic_pairs.ma".
-include "formal_topology/o-saturations.ma".
-
-definition A : ∀b:OBP. unary_morphism1 (Oform b) (Oform b).
-intros; constructor 1;
- [ apply (λx.□⎽b (Ext⎽b x));
- | intros; apply (†(†e));]
-qed.
-
-lemma down_p : ∀S:SET1.∀I:SET.∀u:S ⇒_1 S.∀c:arrows2 SET1 I S.∀a:I.∀a':I.a =_1 a'→u (c a) =_1 u (c a').
-intros; apply (†(†e));
-qed.
-
-record Oconcrete_space : Type2 ≝
- { Obp:> OBP;
- (*distr : is_distributive (form bp);*)
- Odownarrow: unary_morphism1 (Oform Obp) (Oform Obp);
- Odownarrow_is_sat: is_o_saturation ? Odownarrow;
- Oconverges: ∀q1,q2.
- (Ext⎽Obp q1 ∧ (Ext⎽Obp q2)) = (Ext⎽Obp ((Odownarrow q1) ∧ (Odownarrow q2)));
- Oall_covered: Ext⎽Obp (oa_one (Oform Obp)) = oa_one (Oconcr Obp);
- Oil2: ∀I:SET.∀p:arrows2 SET1 I (Oform Obp).
- Odownarrow (∨ { x ∈ I | Odownarrow (p x) | down_p ???? }) =
- ∨ { x ∈ I | Odownarrow (p x) | down_p ???? };
- Oil1: ∀q.Odownarrow (A ? q) = A ? q
- }.
-
-interpretation "o-concrete space downarrow" 'downarrow x =
- (fun11 ?? (Odownarrow ?) x).
-
-definition Obinary_downarrow :
- ∀C:Oconcrete_space.binary_morphism1 (Oform C) (Oform C) (Oform C).
-intros; constructor 1;
-[ intros; apply (↓ c ∧ ↓ c1);
-| intros;
- alias symbol "prop2" = "prop21".
- alias symbol "prop1" = "prop11".
- apply ((†e)‡(†e1));]
-qed.
-
-interpretation "concrete_space binary ↓" 'fintersects a b = (fun21 ? ? ? (Obinary_downarrow ?) a b).
-
-record Oconvergent_relation_pair (CS1,CS2: Oconcrete_space) : Type2 ≝
- { Orp:> arrows2 ? CS1 CS2;
- Orespects_converges:
- ∀b,c. eq1 ? (Orp\sub\c⎻ (Ext⎽CS2 (b ↓ c))) (Ext⎽CS1 (Orp\sub\f⎻ b ↓ Orp\sub\f⎻ c));
- Orespects_all_covered:
- eq1 ? (Orp\sub\c⎻ (Ext⎽CS2 (oa_one (Oform CS2))))
- (Ext⎽CS1 (oa_one (Oform CS1)))
- }.
-
-definition Oconvergent_relation_space_setoid: Oconcrete_space → Oconcrete_space → setoid2.
- intros (c c1);
- constructor 1;
- [ apply (Oconvergent_relation_pair c c1)
- | constructor 1;
- [ intros (c2 c3);
- apply (Orelation_pair_equality c c1 c2 c3);
- | intros 1; apply refl2;
- | intros 2; apply sym2;
- | intros 3; apply trans2]]
-qed.
-
-definition Oconvergent_relation_space_of_Oconvergent_relation_space_setoid:
- ∀CS1,CS2.carr2 (Oconvergent_relation_space_setoid CS1 CS2) →
- Oconvergent_relation_pair CS1 CS2 ≝ λP,Q,c.c.
-coercion Oconvergent_relation_space_of_Oconvergent_relation_space_setoid.
-
-definition Oconvergent_relation_space_composition:
- ∀o1,o2,o3: Oconcrete_space.
- binary_morphism2
- (Oconvergent_relation_space_setoid o1 o2)
- (Oconvergent_relation_space_setoid o2 o3)
- (Oconvergent_relation_space_setoid o1 o3).
- intros; constructor 1;
- [ intros; whd in t t1 ⊢ %;
- constructor 1;
- [ apply (c1 ∘ c);
- | intros;
- change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (b↓c2))));
- alias symbol "trans" = "trans1".
- apply (.= († (Orespects_converges : ?)));
- apply (Orespects_converges ?? c (c1\sub\f⎻ b) (c1\sub\f⎻ c2));
- | change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (oa_one (Oform o3)))));
- apply (.= (†(Orespects_all_covered :?)));
- apply rule (Orespects_all_covered ?? c);]
- | intros;
- change with (b ∘ a = b' ∘ a');
- change in e with (Orp ?? a = Orp ?? a');
- change in e1 with (Orp ?? b = Orp ?? b');
- apply (e‡e1);]
-qed.
-
-definition OCSPA: category2.
- constructor 1;
- [ apply Oconcrete_space
- | apply Oconvergent_relation_space_setoid
- | intro; constructor 1;
- [ apply id2
- | intros; apply refl1;
- | apply refl1]
- | apply Oconvergent_relation_space_composition
- | intros; simplify; whd in a12 a23 a34;
- change with (a34 ∘ (a23 ∘ a12) = (a34 ∘ a23) ∘ a12);
- apply rule ASSOC;
- | intros; simplify;
- change with (a ∘ id2 OBP o1 = a);
- apply (id_neutral_right2 : ?);
- | intros; simplify;
- change with (id2 ? o2 ∘ a = a);
- apply (id_neutral_left2 : ?);]
-qed.
-
-definition Oconcrete_space_of_OCSPA : objs2 OCSPA → Oconcrete_space ≝ λx.x.
-coercion Oconcrete_space_of_OCSPA.
-
-definition Oconvergent_relation_space_setoid_of_arrows2_OCSPA :
- ∀P,Q. arrows2 OCSPA P Q → Oconvergent_relation_space_setoid P Q ≝ λP,Q,x.x.
-coercion Oconvergent_relation_space_setoid_of_arrows2_OCSPA.
-
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/o-basic_topologies.ma".
-
-(*
-(*
-definition downarrow: ∀S:BTop. unary_morphism (Ω \sup S) (Ω \sup S).
- intros; constructor 1;
- [ apply (λU:Ω \sup S.{a | ∃b:carrbt S. b ∈ U ∧ a ∈ A ? (singleton ? b)});
- intros; simplify; split; intro; cases H1; cases x; exists [1,3: apply w]
- split; try assumption; [ apply (. H‡#) | apply (. H \sup -1‡#) ] assumption
- | intros; split; intros 2; cases f; exists; [1,3: apply w] cases x; split;
- try assumption; [ apply (. #‡H) | apply (. #‡H \sup -1)] assumption]
-qed.
-
-interpretation "downarrow" 'downarrow a = (fun_1 ?? (downarrow ?) a).
-
-definition ffintersects: ∀S:BTop. binary_morphism1 (Ω \sup S) (Ω \sup S) (Ω \sup S).
- intros; constructor 1;
- [ apply (λU,V. ↓U ∩ ↓V);
- | intros; apply (.= (†H)‡(†H1)); apply refl1]
-qed.
-
-interpretation "ffintersects" 'fintersects U V = (fun1 ??? (ffintersects ?) U V).
-*)
-
-record formal_topology: Type ≝
- { bt:> OBTop;
- converges: ∀U,V: bt. oA bt (U ↓ V) = (oA ? U ∧ oA ? V)
- }.
-
-(*
-
-definition ffintersects': ∀S:BTop. binary_morphism1 S S (Ω \sup S).
- intros; constructor 1;
- [ apply (λb,c:S. (singleton S b) ↓ (singleton S c));
- | intros; apply (.= (†H)‡(†H1)); apply refl1]
-qed.
-
-interpretation "ffintersects'" 'fintersects U V = (fun1 ??? (ffintersects' ?) U V).
-*)
-record formal_map (S,T: formal_topology) : Type ≝
- { cr:> continuous_relation_setoid S T;
- C1: ∀b,c. extS ?? cr (b ↓ c) = ext ?? cr b ↓ ext ?? cr c;
- C2: extS ?? cr T = S
- }.
-
-definition formal_map_setoid: formal_topology → formal_topology → setoid1.
- intros (S T); constructor 1;
- [ apply (formal_map S T);
- | constructor 1;
- [ apply (λf,f1: formal_map S T.f=f1);
- | simplify; intros 1; apply refl1
- | simplify; intros 2; apply sym1
- | simplify; intros 3; apply trans1]]
-qed.
-
-axiom C1':
- ∀S,T: formal_topology.∀f:formal_map_setoid S T.∀U,V: Ω \sup T.
- extS ?? f (U ↓ V) = extS ?? f U ↓ extS ?? f V.
-
-definition formal_map_composition:
- ∀o1,o2,o3: formal_topology.
- binary_morphism1
- (formal_map_setoid o1 o2)
- (formal_map_setoid o2 o3)
- (formal_map_setoid o1 o3).
- intros; constructor 1;
- [ intros; whd in c c1; constructor 1;
- [ apply (comp1 BTop ??? c c1);
- | intros;
- apply (.= (extS_com ??? c c1 ?));
- apply (.= †(C1 ?????));
- apply (.= (C1' ?????));
- apply (.= ((†((extS_singleton ????) \sup -1))‡(†((extS_singleton ????) \sup -1))));
- apply (.= (extS_com ??????)\sup -1 ‡ (extS_com ??????) \sup -1);
- apply (.= (extS_singleton ????)‡(extS_singleton ????));
- apply refl1;
- | apply (.= (extS_com ??? c c1 ?));
- apply (.= (†(C2 ???)));
- apply (.= (C2 ???));
- apply refl1;]
- | intros; simplify;
- change with (comp1 BTop ??? a b = comp1 BTop ??? a' b');
- apply prop1; assumption]
-qed.
-*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/o-algebra.ma".
-
-definition is_o_saturation: ∀C:OA. C ⇒_1 C → CProp1 ≝
- λC:OA.λA:C ⇒_1 C.∀U,V. (U ≤ A V) =_1 (A U ≤ A V).
-
-definition is_o_reduction: ∀C:OA. C ⇒_1 C → CProp1 ≝
- λC:OA.λJ:C ⇒_1 C.∀U,V. (J U ≤ V) =_1 (J U ≤ J V).
-
-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 o_saturation_monotone: ∀C:OA.∀A:C ⇒_1 C. 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 o_saturation_expansive ]
- assumption.
-qed.
-
-theorem o_saturation_idempotent: ∀C:OA.∀A:C ⇒_1 C. is_o_saturation C A → ∀U. A (A U) =_1 A U.
- intros; apply (oa_leq_antisym C);
- [ apply (if ?? (i (A U) U)); apply (oa_leq_refl C).
- | apply o_saturation_expansive; assumption]
-qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/basic_pairs_to_o-basic_pairs.ma".
-include "formal_topology/apply_functor.ma".
-
-definition rOBP ≝ Apply (category2_of_category1 BP) OBP BP_to_OBP.
-
-include "formal_topology/o-basic_pairs_to_o-basic_topologies.ma".
-
-lemma category2_of_category1_respects_comp_r:
- ∀C:category1.∀o1,o2,o3:C.
- ∀f:arrows1 ? o1 o2.∀g:arrows1 ? o2 o3.
- (comp1 ???? f g) =_\ID (comp2 (category2_of_category1 C) o1 o2 o3 f g).
- intros; constructor 1;
-qed.
-
-lemma category2_of_category1_respects_comp:
- ∀C:category1.∀o1,o2,o3:C.
- ∀f:arrows1 ? o1 o2.∀g:arrows1 ? o2 o3.
- (comp2 (category2_of_category1 C) o1 o2 o3 f g) =_\ID (comp1 ???? f g).
- intros; constructor 1;
-qed.
-
-lemma POW_full':
- ∀S,T:REL.∀f:arrows2 SET1 (POW S) (POW T).
- arrows1 REL S T.
- intros;
- constructor 1; constructor 1;
- [ intros (x y); apply (y ∈ c {(x)});
- | apply hide; intros; unfold FunClass_1_OF_Ocontinuous_relation;
- apply (e1ࠠe); ]
-qed.
-
-(*
-lemma POW_full_image:
- ∀S,T:REL.∀f:arrows2 SET1 (POW S) (POW T).
- exT22 ? (λg:arrows1 REL S T.or_f ?? (map_arrows2 ?? POW S T g) = f).
- intros; letin g ≝ (? : carr1 (arrows1 REL S T)); [
- constructor 1; constructor 1;
- [ intros (x y); apply (y ∈ f {(x)});
- | apply hide; intros; unfold FunClass_1_OF_Ocontinuous_relation;
- apply (e1ࠠe); ]]
-exists [apply g]
-intro; split; intro; simplify; intro;
-[ whd in f1; change in f1:(? ? (λ_:?.? ? ? ? ? % ?)) with (a1 ∈ f {(x)});
- cases f1; cases x; clear f1 x; change with (a1 ∈ f a);
- lapply (f_image_monotone ?? (map_arrows2 ?? POW S T g) (singleton ? w) a ? a1);
- [2: whd in Hletin;
- change in Hletin:(? ? (λ_:?.? ? ? ? ? % ?))
- with (a1 ∈ f {(x)}); cases Hletin; cases x;
- [ intros 2; change in f3 with (eq1 ? w a2); change with (a2 ∈ a);
- apply (. f3^-1‡#); assumption;
- | assumption; ]
-
-
-
- lapply (. (or_prop3 ?? (map_arrows2 ?? POW S T g) (singleton ? a1) a)^-1);
- [ whd in Hletin:(? ? ? ? ? ? %);
- change in Hletin:(? ? ? ? ? ? (? ? (? ? ? (λ_:?.? ? (λ_:?.? ? ? ? ? % ?)) ?)))
- with (y ∈ f {(x)});
- cases Hletin; cases x1; cases x2;
-
- [ cases Hletin; change in x with (eq1 ? a1 w1); apply (. x‡#); assumption;
- | exists; [apply w] assumption ]
-
-
- clear g;
- cases f1; cases x; simplify in f2; change with (a1 ∈ (f a));
- lapply depth=0 (let x ≝ POW in or_prop3 (POW S) (POW T) (map_arrows2 ?? POW S T g));
- lapply (Hletin {(w)} {(a1)}).
- lapply (if ?? Hletin1); [2: clear Hletin Hletin1;
- exists; [apply a1] [whd; exists[apply w] split; [assumption;|change with (w = w); apply rule #]]
- change with (a1=a1); apply rule #;]
- clear Hletin Hletin1; cases Hletin2; whd in x2;
-qed.
-*)
-lemma curry: ∀A,B,C.(A × B ⇒_1 C) → A → (B ⇒_1 C).
- intros;
- constructor 1;
- [ apply (b c);
- | intros; apply (#‡e); ]
-qed.
-
-notation < "F x" left associative with precedence 60 for @{'map_arrows $F $x}.
-interpretation "map arrows 2" 'map_arrows F x = (fun12 ? ? (map_arrows2 ? ? F ? ?) x).
-
-definition preserve_sup : ∀S,T.∀ f:Ω^S ⇒_1 Ω^T. CProp1.
-intros (S T f); apply (∀X:Ω \sup S. (f X) =_1 ?);
-constructor 1; constructor 1;
-[ intro y; alias symbol "singl" = "singleton". alias symbol "and" = "and_morphism".
- apply (∃x:S. x ∈ X ∧ y ∈ f {(x)});
-| intros (a b H); split; intro E; cases E; clear E; exists; [1,3:apply w]
- [ apply (. #‡(H^-1‡#)); | apply (. #‡(H‡#));] assumption]
-qed.
-
-alias symbol "singl" = "singleton".
-lemma eq_cones_to_eq_rel:
- ∀S,T. ∀f,g: arrows1 REL S T.
- (∀x. curry ??? (image ??) f {(x)} = curry ??? (image ??) g {(x)}) → f = g.
-intros; intros 2 (a b); split; intro;
-[ cases (f1 a); lapply depth=0 (s b); clear s s1;
- lapply (Hletin); clear Hletin;
- [ cases Hletin1; cases x; change in f4 with (a = w);
- change with (a ♮g b); apply (. f4‡#); assumption;
- | exists; [apply a] split; [ assumption | change with (a=a); apply rule #;]]
-| cases (f1 a); lapply depth=0 (s1 b); clear s s1;
- lapply (Hletin); clear Hletin;
- [ cases Hletin1; cases x; change in f4 with (a = w);
- change with (a ♮f b); apply (. f4‡#); assumption;
- | exists; [apply a] split; [ assumption | change with (a=a); apply rule #;]]]
-qed.
-
-variant eq_cones_to_eq_rel':
- ∀S,T. ∀f,g: arrows1 REL S T.
- (∀x:S. or_f ?? (map_arrows2 ?? POW S T f) {(x)} = or_f ?? (map_arrows2 ?? POW S T g) {(x)}) →
- f = g
-≝ eq_cones_to_eq_rel.
-
-(*
-lemma rOR_full :
- ∀s,t:rOBP.∀f:arrows2 OBTop (OR (ℱ_2 s)) (OR (ℱ_2 t)).
- exT22 ? (λg:arrows2 rOBP s t.
- map_arrows2 ?? OR ?? (ℳ_2 g) = f).
-intros 2 (s t); cases s (s_2 s_1 s_eq); clear s;
-change in match (F2 ??? (mk_Fo ??????)) with s_2;
-cases s_eq; clear s_eq s_2;
-letin s1 ≝ (BP_to_OBP s_1); change in match (BP_to_OBP s_1) with s1;
-cases t (t_2 t_1 t_eq); clear t;
-change in match (F2 ??? (mk_Fo ??????)) with t_2;
-cases t_eq; clear t_eq t_2;
-letin t1 ≝ (BP_to_OBP t_1); change in match (BP_to_OBP t_1) with t1;
-whd in ⊢ (%→?); whd in ⊢ (? (? ? ? ? %) (? ? ? ? %)→?);
-intro; whd in s_1 t_1;
-letin R ≝ (? : (carr2 (arrows2 (category2_of_category1 BP) s_1 t_1)));
-[2:
- exists;
- [ constructor 1;
- [2: simplify; apply R;
- | simplify; apply (fun12 ?? (map_arrows2 ?? BP_to_OBP s_1 t_1)); apply R;
- | simplify; apply rule #; ]]
- simplify;
-|1: constructor 1;
- [2: apply (pi1exT22 ?? (POW_full (form s_1) (form t_1) f));
- |1: letin u ≝ (or_f_star ?? (map_arrows2 ?? POW (concr t_1) (form t_1) (⊩ \sub t_1)));
- letin r ≝ (u ∘ (or_f ?? f));
- letin xxx ≝ (or_f ?? (map_arrows2 ?? POW (concr s_1) (form s_1) (⊩ \sub s_1)));
- letin r' ≝ (r ∘ xxx); clearbody r';
- apply (POW_full' (concr s_1) (concr t_1) r');
- | simplify in ⊢ (? ? ? (? ? ? ? ? % ?) ?);
- apply eq_cones_to_eq_rel'; intro;
- apply
- (cic:/matita/logic/equality/eq_elim_r''.con ?????
- (category2_of_category1_respects_comp_r : ?));
- apply rule (.= (#‡#));
- apply (.= (respects_comp2 ?? POW (concr s_1) (concr t_1) (form t_1) ? (⊩\sub t_1))‡#);
- apply sym2;
- apply (.= (respects_comp2 ?? POW (concr s_1) (form s_1) (form t_1) (⊩\sub s_1) (pi1exT22 ?? (POW_full (form s_1) (form t_1) (Ocont_rel ?? f)))));
- apply (let H ≝(\snd (POW_full (form s_1) (form t_1) (Ocont_rel ?? f))) in .= #‡H);
- apply sym2;
- ]
-
-STOP;
-
-
-(* Todo: rename BTop → OBTop *)
-
-(* scrivo gli statement qua cosi' verra' un conflitto :-)
-
-1. definire il funtore OR
-2. dimostrare che ORel e' faithful
-
-3. Definire la funzione
- Apply:
- ∀C1,C2: CAT2. F: arrows3 CAT2 C1 C2 → CAT2
- ≝
- constructor 1;
- [ gli oggetti sono gli oggetti di C1 mappati da F
- | i morfismi i morfismi di C1 mappati da F
- | ....
- ]
-
- E : objs CATS === Σx.∃y. F y = x
-
- Quindi (Apply C1 C2 F) (che usando da ora in avanti una coercion
- scrivero' (F C1) ) e' l'immagine di C1 tramite F ed e'
- una sottocategoria di C2 (qualcosa da dimostare qui??? vedi sotto
- al punto 5)
-
-4. Definire rOBP (le OBP rappresentabili) come (BP_to_OBP BP)
- [Si puo' fare lo stesso per le OA: rOA ≝ Rel_to_OA REL ]
-
-5. Dimostrare che OR (il funtore faithful da OBP a OBTop) e' full
- quando applicato a rOBP.
- Nota: puo' darsi che faccia storie ad accettare lo statement.
- Infatti rOBP e' (BP_to_OBP BP) ed e' "una sottocategoria di OBP"
- e OR va da OBP a OBTop. Non so se tipa subito o se devi dare
- una "proiezione" da rOBP a OBP.
-
-6. Definire rOBTop come (OBP_to_OBTop rOBP).
-
-7. Per composizione si ha un funtore full and faithful da BP a rOBTop:
- basta prendere (OR ∘ BP_to_OBP).
-
-8. Dimostrare (banale: quasi tutti i campi sono per conversione) che
- esiste un funtore da rOBTop a BTop. Dimostrare che tale funtore e'
- faithful e full (banale: tutta conversione).
-
-9. Per composizione si ha un funtore full and faithful da BP a BTop.
-
-10. Dimostrare che i seguenti funtori sono anche isomorphism-dense
- (http://planetmath.org/encyclopedia/DenseFunctor.html):
-
- BP_to_OBP
- OBP_to_OBTop quando applicato alle rOBP
- OBTop_to_BTop quando applicato alle rOBTop
-
- Concludere per composizione che anche il funtore da BP a BTop e'
- isomorphism-dense.
-
-====== Da qui in avanti non e' "necessario" nulla:
-
-== altre cose mancanti
-
-11. Dimostrare che le r* e le * orrizzontali
- sono isomorfe dando il funtore da r* a * e dimostrando che componendo i
- due funtori ottengo l'identita'
-
-12. La definizione di r* fa schifo: in pratica dici solo come ottieni
- qualcosa, ma non come lo caratterizzeresti. Ora un teorema carino
- e' che una a* (e.g. una aOBP) e' sempre una rOBP dove "a" sta per
- atomic. Dimostrarlo per tutte le r*.
-
-== categorish/future works
-
-13. definire astrattamente la FG-completion e usare quella per
- ottenere le BP da Rel e le OBP da OA.
-
-14. indebolire le OA, generalizzare le costruzioni, etc. come detto
- con Giovanni
-
-*)
-
-*)
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/subsets.ma".
-
-record binary_relation (A,B: SET) : Type1 ≝
- { satisfy:> binary_morphism1 A B CPROP }.
-
-notation < "hvbox (x \nbsp \natur term 90 r \nbsp y)" with precedence 45 for @{'satisfy $r $x $y}.
-notation > "hvbox (x \natur term 90 r y)" with precedence 45 for @{'satisfy $r $x $y}.
-interpretation "relation applied" 'satisfy r x y = (fun21 ??? (satisfy ?? r) x y).
-
-definition binary_relation_setoid: SET → SET → setoid1.
- intros (A B);
- constructor 1;
- [ apply (binary_relation A B)
- | constructor 1;
- [ apply (λA,B.λr,r': binary_relation A B. ∀x,y. r x y ↔ r' x y)
- | simplify; intros 3; split; intro; assumption
- | simplify; intros 5; split; intro;
- [ apply (fi ?? (f ??)) | apply (if ?? (f ??))] assumption
- | simplify; intros 7; split; intro;
- [ apply (if ?? (f1 ??)) | apply (fi ?? (f ??)) ]
- [ apply (if ?? (f ??)) | apply (fi ?? (f1 ??)) ]
- assumption]]
-qed.
-
-definition binary_relation_of_binary_relation_setoid :
- ∀A,B.binary_relation_setoid A B → binary_relation A B ≝ λA,B,c.c.
-coercion binary_relation_of_binary_relation_setoid.
-
-definition composition:
- ∀A,B,C.
- (binary_relation_setoid A B) × (binary_relation_setoid B C) ⇒_1 (binary_relation_setoid A C).
- intros;
- constructor 1;
- [ intros (R12 R23);
- constructor 1;
- constructor 1;
- [ apply (λs1:A.λs3:C.∃s2:B. s1 ♮R12 s2 ∧ s2 ♮R23 s3);
- | intros;
- split; intro; cases e2 (w H3); clear e2; exists; [1,3: apply w ]
- [ apply (. (e^-1‡#)‡(#‡e1^-1)); assumption
- | apply (. (e‡#)‡(#‡e1)); assumption]]
- | intros 8; split; intro H2; simplify in H2 ⊢ %;
- cases H2 (w H3); clear H2; exists [1,3: apply w] cases H3 (H2 H4); clear H3;
- [ lapply (if ?? (e x w) H2) | lapply (fi ?? (e x w) H2) ]
- [ lapply (if ?? (e1 w y) H4)| lapply (fi ?? (e1 w y) H4) ]
- exists; try assumption;
- split; assumption]
-qed.
-
-definition REL: category1.
- constructor 1;
- [ apply setoid
- | intros (T T1); apply (binary_relation_setoid T T1)
- | intros; constructor 1;
- constructor 1; unfold setoid1_of_setoid; simplify;
- [ (* changes required to avoid universe inconsistency *)
- change with (carr o → carr o → CProp); intros; apply (eq ? c c1)
- | intros; split; intro; change in a a' b b' with (carr o);
- change in e with (eq ? a a'); change in e1 with (eq ? b b');
- [ apply (.= (e ^ -1));
- apply (.= e2);
- apply e1
- | apply (.= e);
- apply (.= e2);
- apply (e1 ^ -1)]]
- | apply composition
- | intros 9;
- split; intro;
- cases f (w H); clear f; cases H; clear H;
- [cases f (w1 H); clear f | cases f1 (w1 H); clear f1]
- cases H; clear H;
- exists; try assumption;
- split; try assumption;
- exists; try assumption;
- split; assumption
- |6,7: intros 5; unfold composition; simplify; split; intro;
- unfold setoid1_of_setoid in x y; simplify in x y;
- [1,3: cases e (w H1); clear e; cases H1; clear H1; unfold;
- [ apply (. (e : eq1 ? x w)‡#); assumption
- | apply (. #‡(e : eq1 ? w y)^-1); assumption]
- |2,4: exists; try assumption; split;
- (* change required to avoid universe inconsistency *)
- change in x with (carr o1); change in y with (carr o2);
- first [apply refl | assumption]]]
-qed.
-
-definition setoid_of_REL : objs1 REL → setoid ≝ λx.x.
-coercion setoid_of_REL.
-
-definition binary_relation_setoid_of_arrow1_REL :
- ∀P,Q. arrows1 REL P Q → binary_relation_setoid P Q ≝ λP,Q,x.x.
-coercion binary_relation_setoid_of_arrow1_REL.
-
-
-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_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 (λs.{x | True});
- intros; simplify; split; intro; assumption.
-qed.
-
-coercion full_subset.
-
-definition comprehension: ∀b:REL. (b ⇒_1. CPROP) → Ω^b.
- apply (λb:REL. λP: b ⇒_1 CPROP. {x | P x});
- intros; simplify;
- apply (.= †e); apply refl1.
-qed.
-
-interpretation "subset comprehension" 'comprehension s p =
- (comprehension s (mk_unary_morphism1 ?? p ?)).
-
-definition ext: ∀X,S:REL. (X ⇒_\r1 S) × S ⇒_1 (Ω^X).
- intros (X S); constructor 1;
- [ apply (λr:X ⇒_\r1 S.λf:S.{x ∈ X | x ♮r f}); intros; simplify; apply (.= (e‡#)); apply refl1
- | intros; simplify; split; intros; simplify;
- [ change with (∀x. x ♮a b → x ♮a' b'); intros;
- apply (. (#‡e1^-1)); whd in e; apply (if ?? (e ??)); assumption
- | change with (∀x. x ♮a' b' → x ♮a b); intros;
- apply (. (#‡e1)); whd in e; apply (fi ?? (e ??));assumption]]
-qed.
-
-definition extS: ∀X,S:REL. ∀r:X ⇒_\r1 S. Ω^S ⇒_1 Ω^X.
- (* ∃ is not yet a morphism apply (λX,S,r,F.{x ∈ X | ∃a. a ∈ F ∧ x ♮r a});*)
- intros (X S r); constructor 1;
- [ intro F; constructor 1; constructor 1;
- [ apply (λx. x ∈ X ∧ ∃a:S. a ∈ F ∧ x ♮r a);
- | intros; split; intro; cases f (H1 H2); clear f; split;
- [ apply (. (e^-1‡#)); assumption
- |3: apply (. (e‡#)); assumption
- |2,4: cases H2 (w H3); exists; [1,3: apply w]
- [ apply (. (#‡(e^-1‡#))); assumption
- | apply (. (#‡(e‡#))); assumption]]]
- | intros; split; simplify; intros; cases f; cases e1; split;
- [1,3: assumption
- |2,4: exists; [1,3: apply w]
- [ apply (. (#‡e^-1)‡#); assumption
- | apply (. (#‡e)‡#); assumption]]]
-qed.
-(*
-lemma equalset_extS_id_X_X: ∀o:REL.∀X.extS ?? (id1 ? o) X = X.
- intros;
- unfold extS; simplify;
- split; simplify;
- [ intros 2; change with (a ∈ X);
- cases f; clear f;
- cases H; clear H;
- cases x; clear x;
- change in f2 with (eq1 ? a w);
- apply (. (f2\sup -1‡#));
- assumption
- | intros 2; change in f with (a ∈ X);
- split;
- [ whd; exact I
- | exists; [ apply a ]
- split;
- [ assumption
- | change with (a = a); apply refl]]]
-qed.
-
-lemma extS_com: ∀o1,o2,o3,c1,c2,S. extS o1 o3 (c2 ∘ c1) S = extS o1 o2 c1 (extS o2 o3 c2 S).
- intros; unfold extS; simplify; split; intros; simplify; intros;
- [ cases f (H1 H2); cases H2 (w H3); clear f H2; split; [assumption]
- cases H3 (H4 H5); cases H5 (w1 H6); clear H3 H5; cases H6 (H7 H8); clear H6;
- exists; [apply w1] split [2: assumption] constructor 1; [assumption]
- exists; [apply w] split; assumption
- | cases f (H1 H2); cases H2 (w H3); clear f H2; split; [assumption]
- cases H3 (H4 H5); cases H4 (w1 H6); clear H3 H4; cases H6 (w2 H7); clear H6;
- cases H7; clear H7; exists; [apply w2] split; [assumption] exists [apply w] split;
- assumption]
-qed.
-*)
-
-(* the same as ⋄ for a basic pair *)
-definition image: ∀U,V:REL. (U ⇒_\r1 V) ⇒_2 (Ω^U ⇒_2 Ω^V).
- intros; constructor 1;
- [ intro r; constructor 1;
- [ apply (λS: Ω^U. {y | ∃x:U. x ♮r y ∧ x ∈ S });
- intros; simplify; split; intro; cases e1; exists [1,3: apply w]
- [ apply (. (#‡e^-1)‡#); assumption
- | apply (. (#‡e)‡#); assumption]
- | intros; split;
- [ intro y; simplify; intro yA; cases yA; exists; [ apply w ];
- apply (. #‡(#‡e^-1)); assumption;
- | intro y; simplify; intro yA; cases yA; exists; [ apply w ];
- apply (. #‡(#‡e)); assumption;]]
- | simplify; intros; intro y; simplify; split; simplify; intros (b H); cases H;
- exists; [1,3: apply w]; cases x; split; try assumption;
- [ apply (if ?? (e ??)); | apply (fi ?? (e ??)); ] assumption;]
-qed.
-
-(* the same as □ for a basic pair *)
-definition minus_star_image: ∀U,V:REL. (U ⇒_\r1 V) ⇒_2 (Ω^U ⇒_2 Ω^V).
- intros; constructor 1; intros;
- [ constructor 1;
- [ apply (λS: Ω^U. {y | ∀x:U. x ♮c y → x ∈ S});
- intros; simplify; split; intros; apply f;
- [ apply (. #‡e); | apply (. #‡e ^ -1)] assumption;
- | intros; split; intro; simplify; intros;
- [ apply (. #‡e^-1);| apply (. #‡e); ] apply f; assumption;]
- | intros; intro; simplify; split; simplify; intros; apply f;
- [ apply (. (e x a2)); assumption | apply (. (e^-1 x a2)); assumption]]
-qed.
-
-(* the same as Rest for a basic pair *)
-definition star_image: ∀U,V:REL. (U ⇒_\r1 V) ⇒_2 (Ω^V ⇒_2 Ω^U).
- intros; constructor 1;
- [ intro r; constructor 1;
- [ apply (λS: Ω \sup V. {x | ∀y:V. x ♮r y → y ∈ S});
- intros; simplify; split; intros; apply f;
- [ apply (. e ‡#);| apply (. e^ -1‡#);] assumption;
- | intros; split; simplify; intros;
- [ apply (. #‡e^-1);| apply (. #‡e); ] apply f; assumption;]
- | intros; intro; simplify; split; simplify; intros; apply f;
- [ apply (. e a2 y); | apply (. e^-1 a2 y)] assumption;]
-qed.
-
-(* the same as Ext for a basic pair *)
-definition minus_image: ∀U,V:REL. (U ⇒_\r1 V) ⇒_2 (Ω^V ⇒_2 Ω^U).
- intros; constructor 1;
- [ intro r; constructor 1;
- [ apply (λS: Ω^V. {x | ∃y:V. x ♮r y ∧ y ∈ S }).
- intros; simplify; split; intros; cases e1; cases x; exists; [1,3: apply w]
- split; try assumption; [ apply (. (e^-1‡#)); | apply (. (e‡#));] assumption;
- | intros; simplify; split; simplify; intros; cases e1; cases x;
- exists [1,3: apply w] split; try assumption;
- [ apply (. (#‡e^-1)); | apply (. (#‡e));] assumption]
- | intros; intro; simplify; split; simplify; intros; cases e1; exists [1,3: apply w]
- cases x; split; try assumption;
- [ apply (. e^-1 a2 w); | apply (. e a2 w)] assumption;]
-qed.
-
-definition foo : ∀o1,o2:REL.carr1 (o1 ⇒_\r1 o2) → carr2 (setoid2_of_setoid1 (o1 ⇒_\r1 o2)) ≝ λo1,o2,x.x.
-
-interpretation "relation f⎻*" 'OR_f_minus_star r = (fun12 ?? (minus_star_image ? ?) (foo ?? r)).
-interpretation "relation f⎻" 'OR_f_minus r = (fun12 ?? (minus_image ? ?) (foo ?? r)).
-interpretation "relation f*" 'OR_f_star r = (fun12 ?? (star_image ? ?) (foo ?? r)).
-
-definition image_coercion: ∀U,V:REL. (U ⇒_\r1 V) → Ω^U ⇒_2 Ω^V.
-intros (U V r Us); apply (image U V r); qed.
-coercion image_coercion.
-
-(* minus_image is the same as ext *)
-
-theorem image_id: ∀o. (id1 REL o : carr2 (Ω^o ⇒_2 Ω^o)) =_1 (id2 SET1 Ω^o).
- intros; unfold image_coercion; unfold image; simplify;
- whd in match (?:carr2 ?);
- intro U; simplify; split; simplify; intros;
- [ change with (a ∈ U);
- cases e; cases x; change in e1 with (w =_1 a); apply (. e1^-1‡#); assumption
- | change in f with (a ∈ U);
- exists; [apply a] split; [ change with (a = a); apply refl1 | assumption]]
-qed.
-
-theorem minus_star_image_id: ∀o:REL.
- ((id1 REL o)⎻* : carr2 (Ω^o ⇒_2 Ω^o)) =_1 (id2 SET1 Ω^o).
- intros; unfold minus_star_image; simplify; intro U; simplify;
- split; simplify; intros;
- [ change with (a ∈ U); apply f; change with (a=a); apply refl1
- | change in f1 with (eq1 ? x a); apply (. f1‡#); apply f]
-qed.
-
-alias symbol "compose" (instance 5) = "category2 composition".
-alias symbol "compose" (instance 4) = "category1 composition".
-theorem image_comp: ∀A,B,C.∀r:B ⇒_\r1 C.∀s:A ⇒_\r1 B.
- ((r ∘ s) : carr2 (Ω^A ⇒_2 Ω^C)) =_1 r ∘ s.
- intros; intro U; split; intro x; (unfold image; unfold SET1; simplify);
- intro H; cases H;
- cases x1; [cases f|cases f1]; exists; [1,3: apply w1] cases x2; split; try assumption;
- exists; try assumption; split; assumption;
-qed.
-
-theorem minus_star_image_comp:
- ∀A,B,C.∀r:B ⇒_\r1 C.∀s:A ⇒_\r1 B.
- minus_star_image A C (r ∘ s) =_1 minus_star_image B C r ∘ (minus_star_image A B s).
- intros; unfold minus_star_image; intro X; simplify; split; simplify; intros;
- [ whd; intros; simplify; whd; intros; apply f; exists; try assumption; split; assumption;
- | cases f1; cases x1; apply f; assumption]
-qed.
-
-
-(*
-(*CSC: unused! *)
-theorem ext_comp:
- ∀o1,o2,o3: REL.
- ∀a: arrows1 ? o1 o2.
- ∀b: arrows1 ? o2 o3.
- ∀x. ext ?? (b∘a) x = extS ?? a (ext ?? b x).
- intros;
- unfold ext; unfold extS; simplify; split; intro; simplify; intros;
- cases f; clear f; split; try assumption;
- [ cases f2; clear f2; cases x1; clear x1; exists; [apply w] split;
- [1: split] assumption;
- | cases H; clear H; cases x1; clear x1; exists [apply w]; split;
- [2: cases f] assumption]
-qed.
-*)
-
-axiom daemon : False.
-
-theorem extS_singleton:
- ∀o1,o2.∀a.∀x.extS o1 o2 a {(x)} = ext o1 o2 a x.
- intros; unfold extS; unfold ext; unfold singleton; simplify;
- split; intros 2; simplify; simplify in f;
- [ cases f; cases e; cases x1; change in f2 with (x =_1 w); apply (. #‡f2); assumption;
- | split; try apply I; exists [apply x] split; try assumption; change with (x = x); apply rule #;] qed.
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/relations.ma".
-include "formal_topology/o-algebra.ma".
-
-definition POW': objs1 SET → OAlgebra.
- intro A; constructor 1;
- [ apply (Ω^A);
- | apply subseteq;
- | apply overlaps;
- | apply big_intersects;
- | apply big_union;
- | apply ({x | True});
- simplify; intros; apply (refl1 ? (eq1 CPROP));
- | apply ({x | False});
- simplify; intros; apply (refl1 ? (eq1 CPROP));
- | intros; whd; intros; assumption
- | intros; whd; split; assumption
- | intros; apply transitive_subseteq_operator; [2: apply f; | skip | assumption]
- | intros; cases f; exists [apply w] assumption
- | intros; split; [ intros 4; apply (f ? f1 i); | intros 3; intro; apply (f i ? f1); ]
- | intros; split;
- [ intros 4; apply f; exists; [apply i] assumption;
- | intros 3; intros; cases f1; apply (f w a x); ]
- | intros 3; cases f;
- | intros 3; constructor 1;
- | intros; cases f; exists; [apply w]
- [ assumption
- | whd; intros; cases i; simplify; assumption]
- | intros; split; intro;
- [ (** screenshot "screen-pow". *) cases f; cases x1; exists [apply w1] exists [apply w] assumption;
- | cases e; cases x; exists; [apply w1] [ assumption | exists; [apply w] assumption]]
- | intros; intros 2; cases (f {(a)} ?);
- [ exists; [apply a] [assumption | change with (a = a); apply refl1;]
- | change in x1 with (a = w); change with (mem A a q); apply (. (x1‡#));
- assumption]]
-qed.
-
-definition powerset_of_POW': ∀A.oa_P (POW' A) → Ω^A ≝ λA,x.x.
-coercion powerset_of_POW'.
-
-definition orelation_of_relation: ∀o1,o2:REL. o1 ⇒_\r1 o2 → (POW' o1) ⇒_\o2 (POW' o2).
- intros;
- constructor 1;
- [ apply rule c;
- | apply rule (c⎻* );
- | apply rule (c* );
- | apply rule (c⎻);
- | intros; split; intro;
- [ intros 2; intros 2; apply (f y); exists[apply a] split; assumption;
- | intros 2; change with (a ∈ q); cases f1; cases x; clear f1 x;
- apply (f w f3); assumption; ]
- | unfold foo; intros; split; intro;
- [ intros 2; intros 2; apply (f x); exists [apply a] split; assumption;
- | intros 2; change with (a ∈ q); cases f1; cases x; apply (f w f3); assumption;]
- | intros; split; unfold foo; unfold image_coercion; simplify; intro; cases f; clear f;
- [ cases x; cases x2; clear x x2; exists; [apply w1]
- [ assumption | exists; [apply w] split; assumption]
- | cases x1; cases x2; clear x1 x2; exists; [apply w1]
- [ exists; [apply w] split; assumption;
- | assumption; ]]]
-qed.
-
-lemma orelation_of_relation_preserves_equality:
- ∀o1,o2:REL.∀t,t': o1 ⇒_\r1 o2.
- t = t' → orelation_of_relation ?? t =_2 orelation_of_relation ?? t'.
- intros; split; unfold orelation_of_relation; unfold foo; simplify;
- change in e with (t =_2 t'); unfold image_coercion; apply (†e);
-qed.
-
-lemma minus_image_id : ∀o:REL.((id1 REL o))⎻ =_1 (id2 SET1 Ω^o).
-unfold foo; intro o; intro; unfold minus_image; simplify; split; simplify; intros;
-[ cases e; cases x; change with (a1 ∈ a); change in f with (a1 =_1 w);
- apply (. f‡#); assumption;
-| change in f with (a1 ∈ a); exists [ apply a1] split; try assumption;
- change with (a1 =_1 a1); apply refl1;]
-qed.
-
-lemma star_image_id : ∀o:REL. ((id1 REL o))* =_1 (id2 SET1 Ω^o).
-unfold foo; intro o; intro; unfold star_image; simplify; split; simplify; intros;
-[ change with (a1 ∈ a); apply f; change with (a1 =_1 a1); apply rule refl1;
-| change in f1 with (a1 =_1 y); apply (. f1^-1‡#); apply f;]
-qed.
-
-lemma orelation_of_relation_preserves_identity:
- ∀o1:REL. orelation_of_relation ?? (id1 ? o1) =_2 id2 OA (POW' o1).
- intros; split;
- (unfold orelation_of_relation; unfold OA; unfold foo; simplify);
- [ apply (minus_star_image_id o1);
- | apply (minus_image_id o1);
- | apply (image_id o1);
- | apply (star_image_id o1) ]
-qed.
-
-(*
- split; whd; intro;
- [ change with ((∀x. x ♮(id1 REL o1) a1→x∈a) → a1 ∈ a); intros;
- apply (f a1); change with (a1 = a1); apply refl1;
- | change with (a1 ∈ a → ∀x. x ♮(id1 REL o1) a1→x∈a); intros;
- change in f1 with (x = a1); apply (. f1‡#); apply f;
- | alias symbol "and" = "and_morphism".
- change with ((∃y:o1.a1 ♮(id1 REL o1) y ∧ y∈a) → a1 ∈ a);
- intro; cases e; clear e; cases x; clear x; change in f with (a1=w);
- apply (. f‡#); apply f1;
- | change with (a1 ∈ a → ∃y:o1.a1 ♮(id1 REL o1) y ∧ y∈a);
- intro; exists; [apply a1]; split; [ change with (a1=a1); apply refl1; | apply f]
- | change with ((∃x:o1.x ♮(id1 REL o1) a1∧x∈a) → a1 ∈ a);
- intro; cases e; clear e; cases x; clear x; change in f with (w=a1);
- apply (. f^-1‡#); apply f1;
- | change with (a1 ∈ a → ∃x:o1.x ♮(id1 REL o1) a1∧x∈a);
- intro; exists; [apply a1]; split; [ change with (a1=a1); apply refl1; | apply f]
- | change with ((∀y.a1 ♮(id1 REL o1) y→y∈a) → a1 ∈ a); intros;
- apply (f a1); change with (a1 = a1); apply refl1;
- | change with (a1 ∈ a → ∀y.a1 ♮(id1 REL o1) y→y∈a); intros;
- change in f1 with (a1 = y); apply (. f1^-1‡#); apply f;]
-qed.
-*)
-
-(* CSC: ???? forse un uncertain mancato *)
-alias symbol "eq" = "setoid2 eq".
-alias symbol "compose" = "category1 composition".
-lemma orelation_of_relation_preserves_composition:
- ∀o1,o2,o3:REL.∀F: o1 ⇒_\r1 o2.∀G: o2 ⇒_\r1 o3.
- orelation_of_relation ?? (G ∘ F) =
- comp2 OA ??? (orelation_of_relation ?? F) (orelation_of_relation ?? G).
- intros; split; intro; split; whd; intro; whd in ⊢ (% → %); intros;
- [ whd; intros; apply f; exists; [ apply x] split; assumption;
- | cases f1; clear f1; cases x1; clear x1; apply (f w); assumption;
- | cases e; cases x; cases f; cases x1; clear e x f x1; exists; [ apply w1 ]
- split; [ assumption | exists; [apply w] split; assumption ]
- | cases e; cases x; cases f1; cases x1; clear e x f1 x1; exists; [apply w1 ]
- split; [ exists; [apply w] split; assumption | assumption ]
- | unfold arrows1_of_ORelation_setoid;
- cases e; cases x; cases f; cases x1; clear e x f x1; exists; [ apply w1 ]
- split; [ assumption | exists; [apply w] split; assumption ]
- | unfold arrows1_of_ORelation_setoid in e;
- cases e; cases x; cases f1; cases x1; clear e x f1 x1; exists; [apply w1 ]
- split; [ exists; [apply w] split; assumption | assumption ]
- | whd; intros; apply f; exists; [ apply y] split; assumption;
- | cases f1; clear f1; cases x; clear x; apply (f w); assumption;]
-qed.
-
-definition POW: carr3 (arrows3 CAT2 (category2_of_category1 REL) OA).
- constructor 1;
- [ apply POW';
- | intros; constructor 1;
- [ apply (orelation_of_relation S T);
- | intros; apply (orelation_of_relation_preserves_equality S T a a' e); ]
- | apply orelation_of_relation_preserves_identity;
- | apply orelation_of_relation_preserves_composition; ]
-qed.
-
-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; simplify; unfold image_coercion in e3; cases (e3 {(x)});
- split; intro; [ lapply (s y); | lapply (s1 y); ]
- [2,4: exists; [1,3:apply x] split; [1,3: assumption |*: change with (x=x); apply rule #]
- |*: 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.
-*)
-
-include "formal_topology/notation.ma".
-
-theorem POW_full: full2 ?? POW.
- intros 3 (S T); exists;
- [ constructor 1; constructor 1;
- [ apply (λx:carr S.λy:carr T. y ∈ f {(x)});
- | apply hide; intros; unfold FunClass_1_OF_carr2; lapply (.= e1‡#);
- [4: apply mem; |6: apply Hletin;|1,2,3,5: skip]
- lapply (#‡prop11 ?? f ?? (†e)); [6: apply Hletin; |*:skip ]]
- | (split; intro; split; simplify);
- [ change with (∀a1.(∀x. a1 ∈ (f {(x):S}) → x ∈ a) → a1 ∈ f⎻* a);
- | change with (∀a1.a1 ∈ f⎻* a → (∀x.a1 ∈ f {(x):S} → x ∈ a));
- | alias symbol "and" (instance 4) = "and_morphism".
-change with (∀a1.(∃y:carr T. y ∈ f {(a1):S} ∧ y ∈ a) → a1 ∈ f⎻ a);
- | alias symbol "and" (instance 2) = "and_morphism".
-change with (∀a1.a1 ∈ f⎻ a → (∃y:carr T.y ∈ f {(a1):S} ∧ y ∈ a));
- | alias symbol "and" (instance 3) = "and_morphism".
-change with (∀a1.(∃x:carr S. a1 ∈ f {(x):S} ∧ x ∈ a) → a1 ∈ f a);
- | change with (∀a1.a1 ∈. f a → (∃x:carr S. a1 ∈ f {(x):S} ∧ x ∈ a));
- | change with (∀a1.(∀y. y ∈ f {(a1):S} → y ∈ a) → a1 ∈ f* a);
- | change with (∀a1.a1 ∈ f* a → (∀y. y ∈ f {(a1):S} → y ∈ a)); ]
- [ intros; apply ((. (or_prop2 ?? f (singleton ? a1) a)^-1) ? a1);
- [ intros 2; apply (f1 a2); change in f2 with (a2 ∈ f⎻ (singleton ? a1));
- lapply (. (or_prop3 ?? f (singleton ? a2) (singleton ? a1)));
- [ cases Hletin; change in x1 with (eq1 ? a1 w);
- apply (. x1‡#); assumption;
- | exists; [apply a2] [change with (a2=a2); apply rule #; | assumption]]
- | change with (a1 = a1); apply rule #; ]
- | intros; apply ((. (or_prop2 ?? f (singleton ? a1) a)) ? x);
- [ intros 2; change in f3 with (eq1 ? a1 a2); change with (a2 ∈ f⎻* a); apply (. f3^-1‡#);
- assumption;
- | lapply (. (or_prop3 ?? f (singleton ? x) (singleton ? a1))^-1);
- [ cases Hletin; change in x1 with (eq1 ? x w);
- change with (x ∈ f⎻ (singleton ? a1)); apply (. x1‡#); assumption;
- | exists; [apply a1] [assumption | change with (a1=a1); apply rule #; ]]]
- | intros; cases e; cases x; clear e x;
- lapply (. (or_prop3 ?? f (singleton ? a1) a)^-1);
- [ cases Hletin; change in x with (eq1 ? a1 w1); apply (. x‡#); assumption;
- | exists; [apply w] assumption ]
- | intros; lapply (. (or_prop3 ?? f (singleton ? a1) a));
- [ cases Hletin; exists; [apply w] split; assumption;
- | exists; [apply a1] [change with (a1=a1); apply rule #; | assumption ]]
- | intros; cases e; cases x; clear e x;
- apply (f_image_monotone ?? f (singleton ? w) a ? a1);
- [ intros 2; change in f3 with (eq1 ? w a2); change with (a2 ∈ a);
- apply (. f3^-1‡#); assumption;
- | assumption; ]
- | intros; lapply (. (or_prop3 ?? f a (singleton ? a1))^-1);
- [ cases Hletin; exists; [apply w] split;
- [ lapply (. (or_prop3 ?? f (singleton ? w) (singleton ? a1)));
- [ cases Hletin1; change in x3 with (eq1 ? a1 w1); apply (. x3‡#); assumption;
- | exists; [apply w] [change with (w=w); apply rule #; | assumption ]]
- | assumption ]
- | exists; [apply a1] [ assumption; | change with (a1=a1); apply rule #;]]
- | intros; apply ((. (or_prop1 ?? f (singleton ? a1) a)^-1) ? a1);
- [ apply f1; | change with (a1=a1); apply rule #; ]
- | intros; apply ((. (or_prop1 ?? f (singleton ? a1) a)) ? y);
- [ intros 2; change in f3 with (eq1 ? a1 a2); change with (a2 ∈ f* a);
- apply (. f3^-1‡#); assumption;
- | assumption ]]]
-qed.
\ No newline at end of file
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/relations.ma".
-
-definition is_saturation: ∀C:REL. Ω^C ⇒_1 Ω^C → CProp1 ≝
- λC:REL.λA:Ω^C ⇒_1 Ω^C. ∀U,V. (U ⊆ A V) =_1 (A U ⊆ A V).
-
-definition is_reduction: ∀C:REL. Ω^C ⇒_1 Ω^C → CProp1 ≝
- λC:REL.λJ: Ω^C ⇒_1 Ω^C. ∀U,V. (J U ⊆ V) =_1 (J U ⊆ J V).
-
-theorem saturation_expansive: ∀C,A. is_saturation C A → ∀U. U ⊆ A U.
- intros; apply (fi ?? (i ??)); apply subseteq_refl.
-qed.
-
-theorem saturation_monotone:
- ∀C,A. is_saturation C A →
- ∀U,V. U ⊆ V → A U ⊆ A V.
- intros; apply (if ?? (i ??)); apply subseteq_trans; [apply V|3: apply saturation_expansive ]
- assumption.
-qed.
-
-theorem saturation_idempotent: ∀C,A. is_saturation C A → ∀U. A (A U) = A U.
- intros; split;
- [ apply (if ?? (i ??)); apply subseteq_refl
- | apply saturation_expansive; assumption]
-qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/saturations.ma".
-include "formal_topology/o-saturations.ma".
-include "formal_topology/relations_to_o-algebra.ma".
-
-(* These are only conversions :-) *)
-
-definition o_operator_of_operator: ∀C:REL. (Ω^C ⇒_1 Ω^C) → ((POW C) ⇒_1 (POW C)) ≝ λC,t.t.
-
-definition is_o_saturation_of_is_saturation:
- ∀C:REL.∀R: Ω^C ⇒_1 Ω^C. is_saturation ? R → is_o_saturation ? (o_operator_of_operator ? R).
-intros (C R i); apply i; qed.
-
-definition is_o_reduction_of_is_reduction:
- ∀C:REL.∀R: Ω^C ⇒_1 Ω^C.is_reduction ? R → is_o_reduction ? (o_operator_of_operator ? R).
-intros (C R i); apply i; qed.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "formal_topology/categories.ma".
-
-record powerset_carrier (A: objs1 SET) : Type1 ≝ { mem_operator: A ⇒_1 CPROP }.
-interpretation "powerset low" 'powerset A = (powerset_carrier A).
-notation "hvbox(a break ∈. b)" non associative with precedence 45 for @{ 'mem_low $a $b }.
-interpretation "memlow" 'mem_low a S = (fun11 ?? (mem_operator ? S) a).
-
-definition subseteq_operator: ∀A: objs1 SET. Ω^A → Ω^A → CProp0 ≝
- λA:objs1 SET.λU,V.∀a:A. a ∈. U → a ∈. V.
-
-theorem transitive_subseteq_operator: ∀A. transitive2 ? (subseteq_operator A).
- intros 6; intros 2;
- apply s1; apply s;
- assumption.
-qed.
-
-definition powerset_setoid1: SET → SET1.
- intros (T);
- constructor 1;
- [ apply (powerset_carrier T)
- | constructor 1;
- [ apply (λU,V. subseteq_operator ? U V ∧ subseteq_operator ? V U)
- | simplify; intros; split; intros 2; assumption
- | simplify; intros (x y H); cases H; split; assumption
- | simplify; intros (x y z H H1); cases H; cases H1; split;
- apply transitive_subseteq_operator; [1,4: apply y ]
- assumption ]]
-qed.
-
-interpretation "powerset" 'powerset A = (powerset_setoid1 A).
-
-interpretation "subset construction" 'subset \eta.x =
- (mk_powerset_carrier ? (mk_unary_morphism1 ? CPROP x ?)).
-
-definition mem: ∀A. A × Ω^A ⇒_1 CPROP.
- intros;
- constructor 1;
- [ apply (λx,S. mem_operator ? S x)
- | intros 5;
- cases b; clear b; cases b'; clear b'; simplify; intros;
- apply (trans1 ????? (prop11 ?? u ?? e));
- cases e1; whd in s s1;
- split; intro;
- [ apply s; assumption
- | apply s1; assumption]]
-qed.
-
-interpretation "mem" 'mem a S = (fun21 ??? (mem ?) a S).
-
-definition subseteq: ∀A. Ω^A × Ω^A ⇒_1 CPROP.
- intros;
- constructor 1;
- [ apply (λU,V. subseteq_operator ? U V)
- | intros;
- cases e; cases e1;
- split; intros 1;
- [ apply (transitive_subseteq_operator ????? s2);
- apply (transitive_subseteq_operator ???? s1 s4)
- | apply (transitive_subseteq_operator ????? s3);
- apply (transitive_subseteq_operator ???? s s4) ]]
-qed.
-
-interpretation "subseteq" 'subseteq U V = (fun21 ??? (subseteq ?) U V).
-
-theorem subseteq_refl: ∀A.∀S:Ω^A.S ⊆ S.
- intros 4; assumption.
-qed.
-
-theorem subseteq_trans: ∀A.∀S1,S2,S3: Ω^A. S1 ⊆ S2 → S2 ⊆ S3 → S1 ⊆ S3.
- intros; apply transitive_subseteq_operator; [apply S2] assumption.
-qed.
-
-definition overlaps: ∀A. Ω^A × Ω^A ⇒_1 CPROP.
- intros;
- constructor 1;
- [ apply (λA:objs1 SET.λU,V:Ω^A.(exT2 ? (λx:A.x ∈ U) (λx:A.x ∈ V) : CProp0))
- | intros;
- constructor 1; intro; cases e2; exists; [1,4: apply w]
- [ apply (. #‡e^-1); assumption
- | apply (. #‡e1^-1); assumption
- | apply (. #‡e); assumption;
- | apply (. #‡e1); assumption]]
-qed.
-
-interpretation "overlaps" 'overlaps U V = (fun21 ??? (overlaps ?) U V).
-
-definition intersects: ∀A. Ω^A × Ω^A ⇒_1 Ω^A.
- intros;
- constructor 1;
- [ apply rule (λU,V. {x | x ∈ U ∧ x ∈ V });
- intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1;
- | intros;
- split; intros 2; simplify in f ⊢ %;
- [ apply (. (#‡e^-1)‡(#‡e1^-1)); assumption
- | apply (. (#‡e)‡(#‡e1)); assumption]]
-qed.
-
-interpretation "intersects" 'intersects U V = (fun21 ??? (intersects ?) U V).
-
-definition union: ∀A. Ω^A × Ω^A ⇒_1 Ω^A.
- intros;
- constructor 1;
- [ apply (λU,V. {x | x ∈ U ∨ x ∈ V });
- intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1
- | intros;
- split; intros 2; simplify in f ⊢ %;
- [ apply (. (#‡e^-1)‡(#‡e1^-1)); assumption
- | apply (. (#‡e)‡(#‡e1)); assumption]]
-qed.
-
-interpretation "union" 'union U V = (fun21 ??? (union ?) U V).
-
-(* qua non riesco a mettere set *)
-definition singleton: ∀A:setoid. A ⇒_1 Ω^A.
- intros; constructor 1;
- [ apply (λa:A.{b | a =_0 b}); unfold setoid1_of_setoid; simplify;
- intros; simplify;
- split; intro;
- apply (.= e1);
- [ apply e | apply (e^-1) ]
- | unfold setoid1_of_setoid; simplify;
- intros; split; intros 2; simplify in f ⊢ %; apply trans;
- [ apply a |4: apply a'] try assumption; apply sym; assumption]
-qed.
-
-interpretation "singleton" 'singl a = (fun11 ?? (singleton ?) a).
-notation > "{ term 19 a : S }" with precedence 90 for @{fun11 ?? (singleton $S) $a}.
-
-definition big_intersects: ∀A:SET.∀I:SET.(I ⇒_2 Ω^A) ⇒_2 Ω^A.
- intros; constructor 1;
- [ intro; whd; whd in I;
- apply ({x | ∀i:I. x ∈ c i});
- simplify; intros; split; intros; [ apply (. (e^-1‡#)); | apply (. e‡#); ]
- apply f;
- | intros; split; intros 2; simplify in f ⊢ %; intro;
- [ apply (. (#‡(e i)^-1)); apply f;
- | apply (. (#‡e i)); apply f]]
-qed.
-
-definition big_union: ∀A:SET.∀I:SET.(I ⇒_2 Ω^A) ⇒_2 Ω^A.
- intros; constructor 1;
- [ intro; whd; whd in A; whd in I;
- apply ({x | ∃i:I. x ∈ c i });
- simplify; intros; split; intros; cases e1; clear e1; exists; [1,3:apply w]
- [ apply (. (e^-1‡#)); | apply (. (e‡#)); ]
- apply x;
- | intros; split; intros 2; simplify in f ⊢ %; cases f; clear f; exists; [1,3:apply w]
- [ apply (. (#‡(e w)^-1)); apply x;
- | apply (. (#‡e w)); apply x]]
-qed.
-
-notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋃) \below (\emsp) term 90 p)"
-non associative with precedence 50 for @{ 'bigcup $p }.
-notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋃) \below (ident i ∈ I) break term 90 p)"
-non associative with precedence 50 for @{ 'bigcup_mk (λ${ident i}:$I.$p) }.
-notation > "hovbox(⋃ f)" non associative with precedence 60 for @{ 'bigcup $f }.
-
-notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋂) \below (\emsp) term 90 p)"
-non associative with precedence 50 for @{ 'bigcap $p }.
-notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋂) \below (ident i ∈ I) break term 90 p)"
-non associative with precedence 50 for @{ 'bigcap_mk (λ${ident i}:$I.$p) }.
-notation > "hovbox(⋂ f)" non associative with precedence 60 for @{ 'bigcap $f }.
-
-interpretation "bigcup" 'bigcup f = (fun12 ?? (big_union ??) f).
-interpretation "bigcap" 'bigcap f = (fun12 ?? (big_intersects ??) f).
-interpretation "bigcup mk" 'bigcup_mk f = (fun12 ?? (big_union ??) (mk_unary_morphism2 ?? f ?)).
-interpretation "bigcap mk" 'bigcap_mk f = (fun12 ?? (big_intersects ??) (mk_unary_morphism2 ?? f ?)).
["⇒"; "➾"; "⇨"; "➡"; "⇉"; "⥤"; "⥰"; ] ;
["⇑"; "⇧"; "⬆"; ] ;
["⇓"; "⇩"; "⬇"; ] ;
- ["⇔"; "⬄"; "⬌"; ] ;
+ ["â\86\94"; "â\87\94"; "â¬\84"; "â¬\8c"; ] ;
["≤"; "≲"; "≼"; "≰"; "≴"; "⋠"; ];
["_" ; "⎽"; "⎼"; "⎻"; "⎺"; ];
["<"; "≺"; "≮"; "⊀"; "〈"; "«"; "❬"; "❮"; "❰"; ] ;