]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma
nasty change in the lexer/parser:
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / relations_to_o-algebra.ma
index e661fbfd082333a36cfb52a27b0517f3c42c33d8..842219280ba10fcfe19eba1af70f08b4f8944197 100644 (file)
@@ -15,7 +15,7 @@
 include "relations.ma".
 include "o-algebra.ma".
 
-definition SUBSETS: objs1 SET → OAlgebra.
+definition POW': objs1 SET → OAlgebra.
  intro A; constructor 1;
   [ apply (Ω \sup A);
   | apply subseteq;
@@ -42,40 +42,43 @@ definition SUBSETS: objs1 SET → OAlgebra.
   | intros; split; intro;
      [ 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 (singleton ? a) ?);
+  | 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 orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (SUBSETS o1) (SUBSETS o2).
+definition powerset_of_POW': ∀A.oa_P (POW' A) → Ω \sup A ≝ λA,x.x.
+coercion powerset_of_POW'.
+
+definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (POW' o1) (POW' o2).
  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 +90,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 +111,8 @@ lemma orelation_of_relation_preserves_equality:
     apply (. #‡(e‡#)); ]
 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).
+ ∀o1:REL. eq2 ? (orelation_of_relation ?? (id1 ? o1)) (id2 OA (POW' 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;
@@ -137,16 +135,13 @@ lemma orelation_of_relation_preserves_identity:
     change in f1 with (a1 = y); apply (. f1^-1‡#); 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 *)
+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) =
-   comp2 OA (SUBSETS o1) (SUBSETS o2) (SUBSETS o3)
+   comp2 OA (POW' o1) (POW' o2) (POW' 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;
@@ -156,10 +151,114 @@ 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 ]
-  | cases e; cases x; cases f; cases x1; clear e x f x1; exists; [ apply w1 ]
+  | 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 ]
-  | cases e; cases x; cases f1; cases x1; clear e x f1 x1; exists; [apply w1 ]
+  | 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:
+ ∀S,T.∀f,g:arrows2 (category2_of_category1 REL) S T.
+  map_arrows2 ?? POW ?? f = map_arrows2 ?? POW ?? g → f=g.
+ intros; unfold POW in e; simplify in e; cases e;
+ unfold orelation_of_relation in e3; simplify in e3; clear e e1 e2 e4;
+ intros 2; cases (e3 {(x)}); 
+ split; intro; [ lapply (s y); | lapply (s1 y); ]
+  [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. binary_morphism1 A B C → A → unary_morphism1 B C.
+intros; constructor 1; [ apply (b c); | intros; apply (#‡e);]
+qed.
+
+(*
+alias symbol "singl" = "singleton".
+alias symbol "eq" = "setoid eq".
+lemma in_singleton_to_eq : ∀A:setoid.∀y,x:A.y ∈ {(x)} → (eq1 A) y x.
+intros; apply sym1; apply f;
+qed.  
+
+lemma eq_to_in_singleton : ∀A:setoid.∀y,x:A.eq1 A y x → y ∈ {(x)}.
+intros; apply (e^-1);
+qed.
+*)
+
+interpretation "lifting singl" 'singl x = 
+ (fun11 ? (objs2 (POW ?))  (singleton ?) x).
+
+theorem POW_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? POW S T g = f).
+ intros; exists;
+  [ constructor 1; constructor 1;
+     [ apply (λx:carr S.λy:carr T. y ∈ f {(x)});
+     | 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 ]]  
+     | whd; split; whd; intro; simplify; unfold map_arrows2; simplify; 
+        [ split;
+           [ change with (∀a1.(∀x. a1 ∈ f (singleton S x) → x ∈ a) → a1 ∈ f⎻* a);
+           | change with (∀a1.a1 ∈ f⎻* a → (∀x.a1 ∈ f (singleton S x) → x ∈ a)); ]
+        | split;
+           [ change with (∀a1.(∃y:carr T. y ∈ f (singleton S a1) ∧ y ∈ a) → a1 ∈ f⎻ a);
+           | change with (∀a1.a1 ∈ f⎻ a → (∃y:carr T.y ∈ f (singleton S a1) ∧ y ∈ a)); ]
+        | split;
+           [ change with (∀a1.(∃x:carr S. a1 ∈ f (singleton S x) ∧ x ∈ a) → a1 ∈ f a);
+           | change with (∀a1.a1 ∈ f a → (∃x:carr S. a1 ∈ f (singleton S x) ∧ x ∈ a)); ]
+        | split;
+           [ change with (∀a1.(∀y. y ∈ f (singleton S a1) → y ∈ a) → a1 ∈ f* a);
+           | change with (∀a1.a1 ∈ f* a → (∀y. y ∈ f (singleton S a1) → 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