]> matita.cs.unibo.it Git - helm.git/commitdiff
SUBSETS_full up to universe inconsistency
authorClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Sun, 18 Jan 2009 02:58:50 +0000 (02:58 +0000)
committerClaudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it>
Sun, 18 Jan 2009 02:58:50 +0000 (02:58 +0000)
helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma
helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs.ma
helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs_to_o-basic_topologies.ma
helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma

index 41f9bfd0e0d1e5765f89e285696057b6a3048119..b17dacbaf8a01ab107d8cd24914f9599899e511b 100644 (file)
@@ -317,4 +317,128 @@ definition ORelation_setoid_of_arrows2_OA:
   ∀P,Q. arrows2 OA P Q → ORelation_setoid P Q ≝ λP,Q,c.c.
 coercion ORelation_setoid_of_arrows2_OA.
 
-prefer coercion Type_OF_objs2.
\ No newline at end of file
+prefer coercion Type_OF_objs2.
+
+(* alias symbol "eq" = "setoid1 eq". *)
+
+(* qui la notazione non va *)
+lemma leq_to_eq_join: ∀S:OA.∀p,q:S. p ≤ q → q = (binary_join ? p q).
+ intros;
+ apply oa_leq_antisym;
+  [ apply oa_density; intros;
+    apply oa_overlap_sym;
+    unfold binary_join; simplify;
+    apply (. (oa_join_split : ?));
+    exists; [ apply false ]
+    apply oa_overlap_sym;
+    assumption
+  | unfold binary_join; simplify;
+    apply (. (oa_join_sup : ?)); intro;
+    cases i; whd in ⊢ (? ? ? ? ? % ?);
+     [ assumption | apply oa_leq_refl ]]
+qed.
+
+lemma overlap_monotone_left: ∀S:OA.∀p,q,r:S. p ≤ q → p >< r → q >< r.
+ intros;
+ apply (. (leq_to_eq_join : ?)‡#);
+  [ apply f;
+  | skip
+  | apply oa_overlap_sym;
+    unfold binary_join; simplify;
+    apply (. (oa_join_split : ?));
+    exists [ apply true ]
+    apply oa_overlap_sym;
+    assumption; ]
+qed.
+
+(* Part of proposition 9.9 *)
+lemma f_minus_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R⎻ p ≤ R⎻ q.
+ intros;
+ apply (. (or_prop2 : ?));
+ apply oa_leq_trans; [2: apply f; | skip | apply (. (or_prop2 : ?)^ -1); apply oa_leq_refl;]
+qed.
+(* Part of proposition 9.9 *)
+lemma f_minus_star_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R⎻* p ≤ R⎻* q.
+ intros;
+ apply (. (or_prop2 : ?)^ -1);
+ apply oa_leq_trans; [3: apply f; | skip | apply (. (or_prop2 : ?)); apply oa_leq_refl;]
+qed.
+
+(* Part of proposition 9.9 *)
+lemma f_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R p ≤ R q.
+ intros;
+ apply (. (or_prop1 : ?));
+ apply oa_leq_trans; [2: apply f; | skip | apply (. (or_prop1 : ?)^ -1); apply oa_leq_refl;]
+qed.
+
+(* Part of proposition 9.9 *)
+lemma f_star_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R* p ≤ R* q.
+ intros;
+ apply (. (or_prop1 : ?)^ -1);
+ apply oa_leq_trans; [3: apply f; | skip | apply (. (or_prop1 : ?)); apply oa_leq_refl;]
+qed.
+
+lemma lemma_10_2_a: ∀S,T.∀R:arrows2 OA S T.∀p. p ≤ R⎻* (R⎻ p).
+ intros;
+ apply (. (or_prop2 : ?)^-1);
+ apply oa_leq_refl.
+qed.
+
+lemma lemma_10_2_b: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻ (R⎻* p) ≤ p.
+ intros;
+ apply (. (or_prop2 : ?));
+ apply oa_leq_refl.
+qed.
+
+lemma lemma_10_2_c: ∀S,T.∀R:arrows2 OA S T.∀p. p ≤ R* (R p).
+ intros;
+ apply (. (or_prop1 : ?)^-1);
+ apply oa_leq_refl.
+qed.
+
+lemma lemma_10_2_d: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* p) ≤ p.
+ intros;
+ apply (. (or_prop1 : ?));
+ apply oa_leq_refl.
+qed.
+
+lemma lemma_10_3_a: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻ (R⎻* (R⎻ p)) = R⎻ p.
+ intros; apply oa_leq_antisym;
+  [ apply lemma_10_2_b;
+  | apply f_minus_image_monotone;
+    apply lemma_10_2_a; ]
+qed.
+
+lemma lemma_10_3_b: ∀S,T.∀R:arrows2 OA S T.∀p. R* (R (R* p)) = R* p.
+ intros; apply oa_leq_antisym;
+  [ apply f_star_image_monotone;
+    apply (lemma_10_2_d ?? R p);
+  | apply lemma_10_2_c; ]
+qed.
+
+lemma lemma_10_3_c: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* (R p)) = R p.
+ intros; apply oa_leq_antisym;
+  [ apply lemma_10_2_d;
+  | apply f_image_monotone;
+    apply (lemma_10_2_c ?? R p); ]
+qed.
+
+lemma lemma_10_3_d: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻* (R⎻ (R⎻* p)) = R⎻* p.
+ intros; apply oa_leq_antisym;
+  [ apply f_minus_star_image_monotone;
+    apply (lemma_10_2_b ?? R p);
+  | apply lemma_10_2_a; ]
+qed.
+
+lemma lemma_10_4_a: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻* (R⎻ (R⎻* (R⎻ p))) = R⎻* (R⎻ p).
+ intros; apply (†(lemma_10_3_a ?? R p));
+qed.
+
+lemma lemma_10_4_b: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* (R (R* p))) = R (R* p).
+intros; unfold in ⊢ (? ? ? % %); apply (†(lemma_10_3_b ?? R p));
+qed.
+
+lemma oa_overlap_sym': ∀o:OA.∀U,V:o. (U >< V) = (V >< U).
+ intros; split; intro; apply oa_overlap_sym; assumption.
+qed.
\ No newline at end of file
index 6517689a15696f000993fded2ecf2076aff73dde..3ac2f62eec9831d6b9402213309f5743e5826030 100644 (file)
@@ -209,4 +209,4 @@ interpretation "Universal pre-image ⊩*" 'rest x = (fun12 _ _ (or_f_star _ _) (
 
 notation "'Ext' \sub b" non associative with precedence 90 for @{'ext $b}.
 notation > "'Ext'⎽term 90 b" non associative with precedence 90 for @{'ext $b}.
-interpretation "Existential pre-image ⊩⎻" 'ext x = (fun12 _ _ (or_f_minus _ _) (rel x)).
+interpretation "Existential pre-image ⊩⎻" 'ext x = (fun12 _ _ (or_f_minus _ _) (rel x)).
\ No newline at end of file
index c66e709dd8c14032673936df02bbabe208a864af..7073900b1a4d29650bb3e288e3451d1d65ce8fd0 100644 (file)
 include "o-basic_pairs.ma".
 include "o-basic_topologies.ma".
 
-alias symbol "eq" = "setoid1 eq".
 
-(* qui la notazione non va *)
-lemma leq_to_eq_join: ∀S:OA.∀p,q:S. p ≤ q → q = (binary_join ? p q).
- intros;
- apply oa_leq_antisym;
-  [ apply oa_density; intros;
-    apply oa_overlap_sym;
-    unfold binary_join; simplify;
-    apply (. (oa_join_split : ?));
-    exists; [ apply false ]
-    apply oa_overlap_sym;
-    assumption
-  | unfold binary_join; simplify;
-    apply (. (oa_join_sup : ?)); intro;
-    cases i; whd in ⊢ (? ? ? ? ? % ?);
-     [ assumption | apply oa_leq_refl ]]
-qed.
-
-lemma overlap_monotone_left: ∀S:OA.∀p,q,r:S. p ≤ q → p >< r → q >< r.
- intros;
- apply (. (leq_to_eq_join : ?)‡#);
-  [ apply f;
-  | skip
-  | apply oa_overlap_sym;
-    unfold binary_join; simplify;
-    apply (. (oa_join_split : ?));
-    exists [ apply true ]
-    apply oa_overlap_sym;
-    assumption; ]
-qed.
-
-(* Part of proposition 9.9 *)
-lemma f_minus_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R⎻ p ≤ R⎻ q.
- intros;
- apply (. (or_prop2 : ?));
- apply oa_leq_trans; [2: apply f; | skip | apply (. (or_prop2 : ?)^ -1); apply oa_leq_refl;]
-qed.
-(* Part of proposition 9.9 *)
-lemma f_minus_star_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R⎻* p ≤ R⎻* q.
- intros;
- apply (. (or_prop2 : ?)^ -1);
- apply oa_leq_trans; [3: apply f; | skip | apply (. (or_prop2 : ?)); apply oa_leq_refl;]
-qed.
-
-(* Part of proposition 9.9 *)
-lemma f_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R p ≤ R q.
- intros;
- apply (. (or_prop1 : ?));
- apply oa_leq_trans; [2: apply f; | skip | apply (. (or_prop1 : ?)^ -1); apply oa_leq_refl;]
-qed.
-
-(* Part of proposition 9.9 *)
-lemma f_star_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R* p ≤ R* q.
- intros;
- apply (. (or_prop1 : ?)^ -1);
- apply oa_leq_trans; [3: apply f; | skip | apply (. (or_prop1 : ?)); apply oa_leq_refl;]
-qed.
-
-lemma lemma_10_2_a: ∀S,T.∀R:arrows2 OA S T.∀p. p ≤ R⎻* (R⎻ p).
- intros;
- apply (. (or_prop2 : ?)^-1);
- apply oa_leq_refl.
-qed.
-
-lemma lemma_10_2_b: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻ (R⎻* p) ≤ p.
- intros;
- apply (. (or_prop2 : ?));
- apply oa_leq_refl.
-qed.
-
-lemma lemma_10_2_c: ∀S,T.∀R:arrows2 OA S T.∀p. p ≤ R* (R p).
- intros;
- apply (. (or_prop1 : ?)^-1);
- apply oa_leq_refl.
-qed.
-
-lemma lemma_10_2_d: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* p) ≤ p.
- intros;
- apply (. (or_prop1 : ?));
- apply oa_leq_refl.
-qed.
-
-lemma lemma_10_3_a: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻ (R⎻* (R⎻ p)) = R⎻ p.
- intros; apply oa_leq_antisym;
-  [ apply lemma_10_2_b;
-  | apply f_minus_image_monotone;
-    apply lemma_10_2_a; ]
-qed.
-
-lemma lemma_10_3_b: ∀S,T.∀R:arrows2 OA S T.∀p. R* (R (R* p)) = R* p.
- intros; apply oa_leq_antisym;
-  [ apply f_star_image_monotone;
-    apply (lemma_10_2_d ?? R p);
-  | apply lemma_10_2_c; ]
-qed.
-
-lemma lemma_10_3_c: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* (R p)) = R p.
- intros; apply oa_leq_antisym;
-  [ apply lemma_10_2_d;
-  | apply f_image_monotone;
-    apply (lemma_10_2_c ?? R p); ]
-qed.
-
-lemma lemma_10_3_d: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻* (R⎻ (R⎻* p)) = R⎻* p.
- intros; apply oa_leq_antisym;
-  [ apply f_minus_star_image_monotone;
-    apply (lemma_10_2_b ?? R p);
-  | apply lemma_10_2_a; ]
-qed.
-
-lemma lemma_10_4_a: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻* (R⎻ (R⎻* (R⎻ p))) = R⎻* (R⎻ p).
- intros; apply (†(lemma_10_3_a ?? R p));
-qed.
-
-lemma lemma_10_4_b: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* (R (R* p))) = R (R* p).
-intros; unfold in ⊢ (? ? ? % %); apply (†(lemma_10_3_b ?? R p));
-qed.
-
-lemma oa_overlap_sym': ∀o:OA.∀U,V:o. (U >< V) = (V >< U).
- intros; split; intro; apply oa_overlap_sym; assumption.
-qed.
 
 (* Qui, per fare le cose per bene, ci serve la nozione di funtore categorico *)
 definition o_basic_topology_of_o_basic_pair: BP → BTop.
index 7f327c4972130fcb143b5aa1eb45f2eafb6bbc60..9be9508e019ea712e3f0bba44eeb4b9b8b7b5ace 100644 (file)
@@ -48,6 +48,9 @@ definition SUBSETS: objs1 SET → OAlgebra.
        assumption]]
 qed.
 
+definition powerset_of_SUBSETS: ∀A.oa_P (SUBSETS A) → Ω \sup A ≝ λA,x.x.
+coercion powerset_of_SUBSETS.
+
 definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (SUBSETS o1) (SUBSETS o2).
  intros;
  constructor 1;
@@ -179,4 +182,65 @@ theorem SUBSETS_faithful:
  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 Hletin1; cases x1; change in f3 with (eq1 ? x w); apply (. f3‡#); assumption;]
-qed.
\ No newline at end of file
+qed.
+
+theorem SUBSETS_full: ∀S,T.∀f.∃g. map_arrows2 ?? SUBSETS' S T g = f.
+ intros; exists;
+  [ constructor 1; constructor 1;
+     [ apply (λx.λy. y ∈ f (singleton S 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. y ∈ f (singleton S a1) ∧ y ∈ a) → a1 ∈ f⎻ a);
+           | change with (∀a1.a1 ∈ f⎻ a → (∃y.y ∈ f (singleton S a1) ∧ y ∈ a)); ]
+        | 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. 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.