]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/formal_topology/relations_to_o-algebra.ma
big mess of notation
[helm.git] / helm / software / matita / library / formal_topology / relations_to_o-algebra.ma
index 93bf2c609694693c0bb14b99fb3627cc6b47d0bd..01c06f02046a97a57df349dc2eddd626c100d1d9 100644 (file)
@@ -51,39 +51,25 @@ qed.
 definition powerset_of_POW': ∀A.oa_P (POW' A) → Ω^A ≝ λA,x.x.
 coercion powerset_of_POW'.
 
+definition foo : ∀o1,o2:REL.carr1 (o1 ⇒_\r1 o2) → carr2 (setoid2_of_setoid1 (o1 ⇒_\r1 o2)) ≝ λo1,o2,x.x.
+
 definition orelation_of_relation: ∀o1,o2:REL. o1 ⇒_\r1 o2 → (POW' o1) ⇒_\o2 (POW' o2).
  intros;
  constructor 1;
-  [ constructor 1; 
-     [ apply (λU.image ?? c U);
-     | intros; apply (#‡e); ]
-  | constructor 1;
-     [ apply (λU.minus_star_image ?? c U);
-     | intros; apply (#‡e); ]
-  | constructor 1;
-     [ apply (λU.star_image ?? c U);
-     | intros; apply (#‡e); ]
-  | constructor 1;
-     [ apply (λU.minus_image ?? c U);
-     | intros; apply (#‡e); ]
-  | intros; split; intro;
-     [ 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 ?? c q);
-       change with (∀a. a ∈ image ?? c p → a ∈ q);
-       intros; cases f1; cases x; clear f1 x; apply (f ? f3); assumption; ]
+  [ apply rule c;
+  | apply rule ((foo ?? c)⎻* ); 
+  | apply rule ((foo ?? c)* );
+  | apply rule ((foo ?? c)⎻);
   | intros; split; intro;
-     [ 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 ?? 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;
+     [ 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]
+        [ assumption | exists; [apply w] split; assumption]
      | cases x1; cases x2; clear x1 x2; exists; [apply w1]
         [ exists; [apply w] split; assumption;
         | assumption; ]]]
@@ -92,29 +78,36 @@ 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; 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^-1‡#));
-  | change with (a1 ∈ minus_star_image ?? t' a → a1 ∈ minus_star_image ?? t a);
-    apply (. #‡(e‡#));
-  | change with (a1 ∈ minus_image ?? t a → a1 ∈ minus_image ?? t' a);
-    apply (. #‡(e ^ -1‡#));
-  | change with (a1 ∈ minus_image ?? t' a → a1 ∈ minus_image ?? t a);
-    apply (. #‡(e‡#));
-  | change with (a1 ∈ image ?? t a → a1 ∈ image ?? t' a);
-    apply (. #‡(e ^ -1‡#));
-  | change with (a1 ∈ image ?? t' a → a1 ∈ image ?? t a);
-    apply (. #‡(e‡#));
-  | change with (a1 ∈ star_image ?? t a → a1 ∈ star_image ?? t' a);
-    apply (. #‡(e ^ -1‡#));
-  | change with (a1 ∈ star_image ?? t' a → a1 ∈ star_image ?? t a);
-    apply (. #‡(e‡#)); ]
+ 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.(foo ?? (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.  (foo ?? (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; intro; split; whd; intro; 
+ 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;
@@ -135,6 +128,7 @@ 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.
+*)
 
 (* CSC: ???? forse un uncertain mancato *)
 alias symbol "eq" = "setoid2 eq".
@@ -173,38 +167,43 @@ 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; cases (e3 {(x)}); 
+ 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)});
-     | intros; unfold FunClass_1_OF_carr2; lapply (.= e1‡#);
+     | 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 ]]  
-     | whd; split; whd; intro; simplify; unfold map_arrows2; simplify; 
-        [ split;
+     | (split; intro; split; simplify);
+     (*
+     whd; split; whd; intro; simplify; unfold map_arrows2; simplify; 
+        [ split;*)
            [ 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)); ]
-        | split;
-           [ change with (∀a1.(∃y:carr T. y ∈ f {(a1):S} ∧ y ∈ a) → a1 ∈ f⎻ a);
-           | change with (∀a1.a1 ∈ f⎻ a → (∃y:carr T.y ∈ f {(a1):S} ∧ y ∈ a)); ]
-        | split;
-           [ 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)); ]
-        | split;
-           [ 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)); ]]
+           | 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)));