X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Frelations_to_o-algebra.ma;h=8bf57da98bfccfb58e524cd2e8219ae34d79965d;hb=05cfeb82d2624860e66941421a937f308d66cf33;hp=b3939f90ba68dfd1f1fa57eed6e8003d322f21e9;hpb=13114a0147a28f8c7359c9c19ee254716eb5f55a;p=helm.git 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 index b3939f90b..8bf57da98 100644 --- 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 @@ -15,9 +15,9 @@ 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 (Ω^A); | apply subseteq; | apply overlaps; | apply big_intersects; @@ -40,18 +40,18 @@ definition SUBSETS: objs1 SET → OAlgebra. [ assumption | whd; intros; cases i; simplify; assumption] | intros; split; intro; - [ cases f; cases x1; exists [apply w1] exists [apply w] assumption; + [ (** screenshot "screen-pow". *) 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 powerset_of_SUBSETS: ∀A.oa_P (SUBSETS A) → Ω \sup A ≝ λA,x.x. -coercion powerset_of_SUBSETS. +definition powerset_of_POW': ∀A.oa_P (POW' A) → Ω^A ≝ λA,x.x. +coercion powerset_of_POW'. -definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (SUBSETS o1) (SUBSETS o2). +definition orelation_of_relation: ∀o1,o2:REL. o1 ⇒_\r1 o2 → (POW' o1) ⇒_\o2 (POW' o2). intros; constructor 1; [ constructor 1; @@ -90,7 +90,8 @@ 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. t = t' → eq2 ? (orelation_of_relation ?? t) (orelation_of_relation ?? t'). + ∀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); @@ -112,7 +113,7 @@ lemma orelation_of_relation_preserves_equality: qed. lemma orelation_of_relation_preserves_identity: - ∀o1:REL. eq2 ? (orelation_of_relation ?? (id1 ? o1)) (id2 OA (SUBSETS o1)). + ∀o1:REL. orelation_of_relation ?? (id1 ? o1) =_2 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; @@ -139,11 +140,9 @@ qed. 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) - ?? (*(orelation_of_relation ?? F) (orelation_of_relation ?? G)*). - [ apply (orelation_of_relation ?? F); | apply (orelation_of_relation ?? G); ] + ∀o1,o2,o3:REL.∀F: o1 ⇒_\r1 o2.∀G: o2 ⇒_\r1 o3. + orelation_of_relation ?? (G ∘ F) = + comp2 OA ??? (orelation_of_relation ?? F) (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; @@ -161,9 +160,9 @@ lemma orelation_of_relation_preserves_composition: | cases f1; clear f1; cases x; clear x; apply (f w); assumption;] qed. -definition SUBSETS': carr3 (arrows3 CAT2 (category2_of_category1 REL) OA). +definition POW: carr3 (arrows3 CAT2 (category2_of_category1 REL) OA). constructor 1; - [ apply SUBSETS; + [ apply POW'; | intros; constructor 1; [ apply (orelation_of_relation S T); | intros; apply (orelation_of_relation_preserves_equality S T a a' e); ] @@ -171,38 +170,42 @@ definition SUBSETS': carr3 (arrows3 CAT2 (category2_of_category1 REL) OA). | apply orelation_of_relation_preserves_composition; ] qed. -theorem SUBSETS_faithful: +theorem POW_faithful: ∀S,T.∀f,g:arrows2 (category2_of_category1 REL) S T. - map_arrows2 ?? SUBSETS' ?? f = map_arrows2 ?? SUBSETS' ?? g → f=g. - intros; unfold SUBSETS' in e; simplify in e; cases e; + POW⎽⇒ f =_2 POW⎽⇒ g → f =_2 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; lapply (e3 (singleton ? x)); cases Hletin; + 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 Hletin1; cases x1; change in f3 with (eq1 ? x w); apply (. f3‡#); assumption;] + |*: cases Hletin; cases x1; change in f3 with (x =_1 w); apply (. f3‡#); assumption;] qed. -theorem SUBSETS_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? SUBSETS' S T g = f). +lemma currify: ∀A,B,C. (A × B ⇒_1 C) → A → (B ⇒_1 C). +intros; constructor 1; [ apply (b c); | intros; apply (#‡e);] +qed. + +theorem POW_full: ∀S,T.∀f: (POW S) ⇒_\o2 (POW T) . exT22 ? (λg. POW⎽⇒ g = f). intros; exists; [ constructor 1; constructor 1; - [ apply (λx:carr S.λy:carr T. y ∈ f (singleton S x)); + [ 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)); ] + [ 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 (singleton S a1) ∧ y ∈ a) → a1 ∈ f⎻ a); - | change with (∀a1.a1 ∈ f⎻ a → (∃y:carr T.y ∈ f (singleton S a1) ∧ y ∈ a)); ] + [ 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 (singleton S x) ∧ x ∈ a) → a1 ∈ f a); - | change with (∀a1.a1 ∈ f a → (∃x:carr S. a1 ∈ f (singleton S x) ∧ x ∈ a)); ] + [ 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 (singleton S a1) → y ∈ a) → a1 ∈ f* a); - | change with (∀a1.a1 ∈ f* a → (∀y. y ∈ f (singleton S a1) → y ∈ 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)));