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=9be9508e019ea712e3f0bba44eeb4b9b8b7b5ace;hb=6b71ae123d3e412d43872b8b7965b7013a970d05;hp=7f327c4972130fcb143b5aa1eb45f2eafb6bbc60;hpb=d87c05ef2d706cc9b69454c191568028134138ea;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 7f327c497..9be9508e0 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 @@ -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.