]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma
the new coercion behaviour (variants + composition with ID) and the new
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / relations_to_o-algebra.ma
index 856a7e0aec012945c32de153c841cae010b7b740..3317c0e64672b220864f8c975903cdfe61163286 100644 (file)
@@ -52,30 +52,30 @@ definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (
  intros;
  constructor 1;
   [ constructor 1; 
-     [ apply (λU.image ?? t U);
+     [ apply (λU.image ?? c U);
      | intros; apply (#‡e); ]
   | constructor 1;
-     [ apply (λU.minus_star_image ?? t U);
+     [ apply (λU.minus_star_image ?? c U);
      | intros; apply (#‡e); ]
   | constructor 1;
-     [ apply (λU.star_image ?? t U);
+     [ apply (λU.star_image ?? c U);
      | intros; apply (#‡e); ]
   | constructor 1;
-     [ apply (λU.minus_image ?? t U);
+     [ apply (λU.minus_image ?? c 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);
+     [ change in f with (∀a. a ∈ image ?? c p → a ∈ q);
+       change with (∀a:o1. a ∈ p → a ∈ star_image ?? c 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);
+     | change in f with (∀a:o1. a ∈ p → a ∈ star_image ?? c q);
+       change with (∀a. a ∈ image ?? c 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);
+     [ change in f with (∀a. a ∈ minus_image ?? c p → a ∈ q);
+       change with (∀a:o2. a ∈ p → a ∈ minus_star_image ?? c 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);
+     | change in f with (∀a:o2. a ∈ p → a ∈ minus_star_image ?? c q);
+       change with (∀a. a ∈ minus_image ?? c 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]
@@ -87,7 +87,7 @@ definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (
 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'.
+ ∀o1,o2:REL.∀t,t': arrows1 ? o1 o2. t = t' → eq2 ? (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);
@@ -108,13 +108,8 @@ lemma orelation_of_relation_preserves_equality:
     apply (. #‡(e‡#)); ]
 qed.
 
-lemma hint: ∀o1,o2:OA. Type_OF_setoid21 (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).
+ ∀o1:REL. eq2 ? (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;
@@ -136,13 +131,10 @@ lemma orelation_of_relation_preserves_identity:
   | 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.
-(*
-lemma hint2: ∀S,T. carr2 (arrows2 OA S T) → Type_OF_setoid21 (arrows2 OA S T).
- intros; apply c;
-qed.
-coercion hint2.
-*)
+
 (* 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: arrows1 ? o1 o2.∀G: arrows1 ? o2 o3.
   orelation_of_relation ?? (G ∘ F) =
@@ -156,10 +148,10 @@ lemma orelation_of_relation_preserves_composition:
     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; 
+  | 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; 
+  | 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;