]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma
1) Some reorganization.
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / relations_to_o-algebra.ma
diff --git a/helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma b/helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma
new file mode 100644 (file)
index 0000000..40db818
--- /dev/null
@@ -0,0 +1,173 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||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".
+include "o-algebra.ma".
+
+definition SUBSETS: objs1 SET → OAlgebra.
+ intro A; constructor 1;
+  [ apply (Ω \sup 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; intros 2; apply (f ? f1 i);
+  | intros; intros 2; apply f;
+    (* senza questa change, universe inconsistency *)
+    whd; change in ⊢ (? ? (λ_:%.?)) with (carr I);
+    exists; [apply i] assumption;
+  | intros 3; cases f;
+  | intros 3; constructor 1;
+  | intros; cases f; exists; [apply w]
+     [ assumption
+     | whd; intros; cases i; simplify; assumption]
+  | intros; split; intro;
+     [ cases f; cases x1;
+       (* senza questa change, universe inconsistency *)
+       change in ⊢ (? ? (λ_:%.?)) with (carr I);
+       exists [apply w1] exists [apply w] assumption;
+     | cases e; cases x; exists; [apply w1]
+        [ assumption
+        | (* senza questa change, universe inconsistency *)
+          whd; change in ⊢ (? ? (λ_:%.?)) with (carr I);
+          exists; [apply w] assumption]]
+  | intros; intros 2; cases (f (singleton ? 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 \sup -1‡#));
+       assumption]]
+qed.
+
+definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (SUBSETS o1) (SUBSETS o2).
+ intros;
+ constructor 1;
+  [ constructor 1; 
+     [ apply (λU.image ?? t U);
+     | intros; apply (#‡e); ]
+  | constructor 1;
+     [ apply (λU.minus_star_image ?? t U);
+     | intros; apply (#‡e); ]
+  | constructor 1;
+     [ apply (λU.star_image ?? t U);
+     | intros; apply (#‡e); ]
+  | constructor 1;
+     [ apply (λU.minus_image ?? t U);
+     | intros; apply (#‡e); ]
+  | intros; split; intro;
+     [ change in f with (∀a. a ∈ image ?? t p → a ∈ q);
+       change with (∀a:o1. a ∈ p → a ∈ star_image ?? t q);
+       intros 4; apply f; exists; [apply a] split; assumption;
+     | change in f with (∀a:o1. a ∈ p → a ∈ star_image ?? t q);
+       change with (∀a. a ∈ image ?? t p → a ∈ q);
+       intros; cases f1; cases x; clear f1 x; apply (f ? f3); assumption; ]
+  | intros; split; intro;
+     [ change in f with (∀a. a ∈ minus_image ?? t p → a ∈ q);
+       change with (∀a:o2. a ∈ p → a ∈ minus_star_image ?? t q);
+       intros 4; apply f; exists; [apply a] split; assumption;
+     | change in f with (∀a:o2. a ∈ p → a ∈ minus_star_image ?? t q);
+       change with (∀a. a ∈ minus_image ?? t p → a ∈ q);
+       intros; cases f1; cases x; clear f1 x; apply (f ? f3); assumption; ]
+  | intros; split; 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': arrows1 ? o1 o2. eq1 ? t t' → orelation_of_relation ?? t = orelation_of_relation ?? t'.
+ intros; split; unfold orelation_of_relation; simplify; intro; split; intro;
+ simplify; whd in o1 o2;
+  [ change with (a1 ∈ minus_star_image ?? t a → a1 ∈ minus_star_image ?? t' a);
+    apply (. #‡(e‡#));
+  | change with (a1 ∈ minus_star_image ?? t' a → a1 ∈ minus_star_image ?? t a);
+    apply (. #‡(e ^ -1‡#));
+  | change with (a1 ∈ minus_image ?? t a → a1 ∈ minus_image ?? t' a);
+    apply (. #‡(e‡#));
+  | change with (a1 ∈ minus_image ?? t' a → a1 ∈ minus_image ?? t a);
+    apply (. #‡(e ^ -1‡#));
+  | change with (a1 ∈ image ?? t a → a1 ∈ image ?? t' a);
+    apply (. #‡(e‡#));
+  | change with (a1 ∈ image ?? t' a → a1 ∈ image ?? t a);
+    apply (. #‡(e ^ -1‡#));
+  | change with (a1 ∈ star_image ?? t a → a1 ∈ star_image ?? t' a);
+    apply (. #‡(e‡#));
+  | change with (a1 ∈ star_image ?? t' a → a1 ∈ star_image ?? t a);
+    apply (. #‡(e ^ -1‡#)); ]
+qed.
+
+lemma hint: ∀o1,o2:OA. Type_OF_setoid2 (arrows2 ? o1 o2) → carr2 (arrows2 OA o1 o2).
+ intros; apply t;
+qed.
+coercion hint.
+
+lemma orelation_of_relation_preserves_identity:
+ ∀o1:REL. orelation_of_relation ?? (id1 ? o1) = id2 OA (SUBSETS o1).
+ intros; split; intro; 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 ^ -1‡#); apply f;
+  | alias symbol "and" = "and_morphism".
+    change with ((∃y: carr 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^-1‡#); apply f1;
+  | change with (a1 ∈ a → ∃y: carr o1.a1 ♮(id1 REL o1) y ∧ y∈a);
+    intro; exists; [apply a1]; split; [ change with (a1=a1); apply refl1; | apply f]
+  | change with ((∃x: carr 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‡#); apply f1;
+  | change with (a1 ∈ a → ∃x: carr 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‡#); apply f;]
+qed.
+
+lemma hint2: ∀S,T. carr2 (arrows2 OA S T) → Type_OF_setoid2 (arrows2 OA S T).
+ intros; apply c;
+qed.
+coercion hint2.
+
+(* CSC: ???? forse un uncertain mancato *)
+lemma orelation_of_relation_preserves_composition:
+ ∀o1,o2,o3:REL.∀F: arrows1 ? o1 o2.∀G: arrows1 ? o2 o3.
+  orelation_of_relation ?? (G ∘ F) =
+   comp2 OA (SUBSETS o1) (SUBSETS o2) (SUBSETS o3)
+    ?? (*(orelation_of_relation ?? F) (orelation_of_relation ?? G)*).
+ [ apply (orelation_of_relation ?? F); | apply (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 ]
+  | 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 ]
+  | whd; intros; apply f; exists; [ apply y] split; assumption;
+  | cases f1; clear f1; cases x; clear x; apply (f w); assumption;]
+qed.