--- /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 "relations.ma".
+
+record basic_pair: Type1 ≝
+ { concr: REL;
+ form: REL;
+ rel: arrows1 ? concr form
+ }.
+
+notation "x ⊩ y" with precedence 45 for @{'Vdash2 $x $y}.
+notation "⊩" with precedence 60 for @{'Vdash}.
+
+interpretation "basic pair relation" 'Vdash2 x y = (rel _ x y).
+interpretation "basic pair relation (non applied)" 'Vdash = (rel _).
+
+alias symbol "eq" = "setoid1 eq".
+alias symbol "compose" = "category1 composition".
+record relation_pair (BP1,BP2: basic_pair): Type1 ≝
+ { concr_rel: arrows1 ? (concr BP1) (concr BP2);
+ form_rel: arrows1 ? (form BP1) (form BP2);
+ commute: ⊩ ∘ concr_rel = form_rel ∘ ⊩
+ }.
+
+notation "hvbox (r \sub \c)" with precedence 90 for @{'concr_rel $r}.
+notation "hvbox (r \sub \f)" with precedence 90 for @{'form_rel $r}.
+
+interpretation "concrete relation" 'concr_rel r = (concr_rel __ r).
+interpretation "formal relation" 'form_rel r = (form_rel __ r).
+
+include "o-basic_pairs.ma".
+
+definition o_basic_pair_of_basic_pair: cic:/matita/formal_topology/basic_pairs/basic_pair.ind#xpointer(1/1) → basic_pair.
+ intro;
+ constructor 1;
+ [ apply (SUBSETS (concr b));
+ | apply (SUBSETS (form b));
+ | constructor 1;
+ ]
+qed.
+
+
+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.
+
+lemma eq_to_eq': ∀o1,o2.∀r,r': relation_pair_setoid o1 o2. r=r' → r \sub\f ∘ ⊩ = r'\sub\f ∘ ⊩.
+ intros 7 (o1 o2 r r' H c1 f2);
+ split; intro H1;
+ [ lapply (fi ?? (commute ?? r c1 f2) H1) as H2;
+ lapply (if ?? (H c1 f2) H2) as H3;
+ apply (if ?? (commute ?? r' c1 f2) H3);
+ | lapply (fi ?? (commute ?? r' c1 f2) H1) as H2;
+ lapply (fi ?? (H c1 f2) H2) as H3;
+ apply (if ?? (commute ?? r c1 f2) H3);
+ ]
+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 \sup -1);]
+qed.
+
+definition relation_pair_composition:
+ ∀o1,o2,o3. binary_morphism1 (relation_pair_setoid o1 o2) (relation_pair_setoid o2 o3) (relation_pair_setoid o1 o3).
+ intros;
+ constructor 1;
+ [ 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;
+ apply (.= ASSOC1);
+ apply (.= #‡H1);
+ apply (.= ASSOC1\sup -1);
+ apply (.= H‡#);
+ apply ASSOC1]
+ | intros;
+ change with (⊩ ∘ (b\sub\c ∘ a\sub\c) = ⊩ ∘ (b'\sub\c ∘ a'\sub\c));
+ change in H with (⊩ ∘ a \sub\c = ⊩ ∘ a' \sub\c);
+ change in H1 with (⊩ ∘ b \sub\c = ⊩ ∘ b' \sub\c);
+ apply (.= ASSOC1);
+ apply (.= #‡H1);
+ apply (.= #‡(commute ?? b'));
+ apply (.= ASSOC1 \sup -1);
+ apply (.= H‡#);
+ apply (.= ASSOC1);
+ apply (.= #‡(commute ?? b')\sup -1);
+ apply (ASSOC1 \sup -1)]
+qed.
+
+definition BP: category1.
+ constructor 1;
+ [ apply basic_pair
+ | apply relation_pair_setoid
+ | apply id_relation_pair
+ | apply relation_pair_composition
+ | intros;
+ change with (⊩ ∘ (a34\sub\c ∘ (a23\sub\c ∘ a12\sub\c)) =
+ ⊩ ∘ ((a34\sub\c ∘ a23\sub\c) ∘ a12\sub\c));
+ apply (ASSOC1‡#);
+ | intros;
+ change with (⊩ ∘ (a\sub\c ∘ (id_relation_pair o1)\sub\c) = ⊩ ∘ a\sub\c);
+ apply ((id_neutral_right1 ????)‡#);
+ | intros;
+ change with (⊩ ∘ ((id_relation_pair o2)\sub\c ∘ a\sub\c) = ⊩ ∘ a\sub\c);
+ apply ((id_neutral_left1 ????)‡#);]
+qed.
+
+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 _)).
first [apply refl | assumption]]]
qed.
+(*
definition full_subset: ∀s:REL. Ω \sup s.
apply (λs.{x | True});
intros; simplify; split; intro; assumption.
cases H7; clear H7; exists; [apply w2] split; [assumption] exists [apply w] split;
assumption]
qed.
-
+*)
+axiom daemon: False.
(* the same as ⋄ for a basic pair *)
-definition image: ∀U,V:REL. binary_morphism1 (arrows1 ? U V) (Ω \sup U) (Ω \sup V).
+definition image: ∀U,V:REL. binary_morphism1 (arrows1 ? U V) (Ω \sup U) ?(*(Ω \sup V)*).
+cases daemon; qed.
intros; constructor 1;
[ apply (λr: arrows1 ? U V.λS: Ω \sup U. {y | ∃x:U. x ♮r y ∧ x ∈ S});
- intros; simplify; split; intro; cases H1; exists [1,3: apply w]
- [ apply (. (#‡H)‡#); assumption
- | apply (. (#‡H \sup -1)‡#); assumption]
- | intros; split; simplify; intros; cases H2; exists [1,3: apply w]
- [ apply (. #‡(#‡H1)); cases x; split; try assumption;
- apply (if ?? (H ??)); assumption
- | apply (. #‡(#‡H1 \sup -1)); cases x; split; try assumption;
- apply (if ?? (H \sup -1 ??)); assumption]]
+ intros; simplify; split; intro; cases H; exists [1,3: apply w]
+ [ apply (. (#‡e)‡#); assumption
+ | apply (. (#‡e ^ -1)‡#); assumption]
+ | intros; split; simplify; intros; cases H; exists [1,3: apply w]
+ [ apply (. #‡(#‡e1)); cases x; split; try assumption;
+ apply (if ?? (e ??)); assumption
+ | apply (. #‡(#‡e1 ^ -1)); cases x; split; try assumption;
+ apply (if ?? (e ^ -1 ??)); assumption]]
qed.
+(*
(* the same as □ for a basic pair *)
definition minus_star_image: ∀U,V:REL. binary_morphism1 (arrows1 ? U V) (Ω \sup U) (Ω \sup V).
intros; constructor 1;
[ cases H; cases x1; change in f2 with (eq1 ? x w); apply (. #‡f2 \sup -1);
assumption
| exists; try assumption; split; try assumption; change with (x = x); apply refl]
+qed.
+*)
+
+include "o-algebra.ma".
+
+definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → ORelation (SUBSETS o1) (SUBSETS o2).
+ intros;
+ constructor 1;
+ [
+ |
+ |
+ |
+ |
+ |
+ |
+ ]
qed.
\ No newline at end of file