X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Fsubsets.ma;h=6351fead91eca85d504ca36b96a93d70a4153db2;hb=c78cbede35ed85575e274864e6b6b9c635c6956d;hp=17128ae6a97942fe7831a06baede3ce05443b1dc;hpb=3d7b244a79a1c57d3355deb2f9a70764cde077b9;p=helm.git diff --git a/helm/software/matita/contribs/formal_topology/overlap/subsets.ma b/helm/software/matita/contribs/formal_topology/overlap/subsets.ma index 17128ae6a..6351fead9 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/subsets.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/subsets.ma @@ -12,13 +12,12 @@ (* *) (**************************************************************************) -include "logic/cprop_connectives.ma". include "categories.ma". -record powerset_carrier (A: SET) : Type1 ≝ { mem_operator: unary_morphism1 A CPROP }. +record powerset_carrier (A: objs1 SET) : Type1 ≝ { mem_operator: unary_morphism1 A CPROP }. -definition subseteq_operator: ∀A: SET. powerset_carrier A → powerset_carrier A → CProp2 ≝ - λA:SET.λU,V.∀a:A. mem_operator ? U a → mem_operator ? V a. +definition subseteq_operator: ∀A: SET. powerset_carrier A → powerset_carrier A → CProp0 ≝ + λA:objs1 SET.λU,V.∀a:A. mem_operator ? U a → mem_operator ? V a. theorem transitive_subseteq_operator: ∀A. transitive2 ? (subseteq_operator A). intros 6; intros 2; @@ -74,6 +73,8 @@ qed. interpretation "subseteq" 'subseteq U V = (fun21 ___ (subseteq _) U V). + + theorem subseteq_refl: ∀A.∀S:Ω \sup A.S ⊆ S. intros 4; assumption. qed. @@ -85,9 +86,10 @@ qed. definition overlaps: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP. intros; constructor 1; - [ apply (λA.λU,V:Ω \sup A.exT2 ? (λx:A.x ∈ U) (λx:A.x ∈ V)) + (* Se metto x al posto di ? ottengo una universe inconsistency *) + [ apply (λA:objs1 SET.λU,V:Ω \sup A.(exT2 ? (λx:A.?(*x*) ∈ U) (λx:A.?(*x*) ∈ V) : CProp0)) | intros; - constructor 1; intro; cases H; exists; [1,4: apply w] + constructor 1; intro; cases e2; exists; [1,4: apply w] [ apply (. #‡e); assumption | apply (. #‡e1); assumption | apply (. #‡(e \sup -1)); assumption; @@ -124,6 +126,7 @@ qed. interpretation "union" 'union U V = (fun21 ___ (union _) U V). +(* qua non riesco a mettere set *) definition singleton: ∀A:setoid. unary_morphism1 A (Ω \sup A). intros; constructor 1; [ apply (λa:A.{b | eq ? a b}); unfold setoid1_of_setoid; simplify; @@ -136,4 +139,77 @@ definition singleton: ∀A:setoid. unary_morphism1 A (Ω \sup A). [ apply a |4: apply a'] try assumption; apply sym; assumption] qed. -interpretation "singleton" 'singl a = (fun11 __ (singleton _) a). \ No newline at end of file +interpretation "singleton" 'singl a = (fun11 __ (singleton _) a). + +definition big_intersects: + ∀A:SET.∀I:SET.unary_morphism2 (setoid1_of_setoid I ⇒ Ω \sup A) (setoid2_of_setoid1 (Ω \sup A)). + intros; constructor 1; + [ intro; whd; whd in I; + apply ({x | ∀i:I. x ∈ t i}); + simplify; intros; split; intros; [ apply (. (e‡#)); | apply (. (e \sup -1‡#)); ] + apply f; + | intros; split; intros 2; simplify in f ⊢ %; intro; + [ apply (. (#‡(e i))); apply f; + | apply (. (#‡(e i)\sup -1)); apply f]] +qed. + +(* senza questo exT "fresco", universe inconsistency *) +inductive exT (A:Type0) (P:A→CProp0) : CProp0 ≝ + ex_introT: ∀w:A. P w → exT A P. + +definition big_union: + ∀A:SET.∀I:SET.unary_morphism2 (setoid1_of_setoid I ⇒ Ω \sup A) (setoid2_of_setoid1 (Ω \sup A)). + intros; constructor 1; + [ intro; whd; whd in A; whd in I; + apply ({x | (*∃i:carr I. x ∈ t i*) exT (carr I) (λi. x ∈ t i)}); + simplify; intros; split; intros; cases e1; clear e1; exists; [1,3:apply w] + [ apply (. (e‡#)); | apply (. (e \sup -1‡#)); ] + apply x; + | intros; split; intros 2; simplify in f ⊢ %; cases f; clear f; exists; [1,3:apply w] + [ apply (. (#‡(e w))); apply x; + | apply (. (#‡(e w)\sup -1)); apply x]] +qed. + +(* incluso prima non funziona piu' nulla *) +include "o-algebra.ma". + +definition SUBSETS: objs1 SET → OAlgebra. + intro A; constructor 1; + [ apply (Ω \sup A); + | apply subseteq; + | apply overlaps; + | apply big_intersects; + | apply big_union; + | apply ({x | True}); + simplify; intros; apply (refl1 ? (eq1 CPROP)); + | apply ({x | False}); + simplify; intros; apply (refl1 ? (eq1 CPROP)); + | intros; whd; intros; assumption + | intros; whd; split; assumption + | intros; apply transitive_subseteq_operator; [2: apply f; | skip | assumption] + | intros; cases f; exists [apply w] assumption + | intros; intros 2; apply (f ? f1 i); + | intros; intros 2; apply f; + (* senza questa change, universe inconsistency *) + whd; change in ⊢ (? ? (λ_:%.?)) with (carr I); + exists; [apply i] assumption; + | intros 3; cases f; + | intros 3; constructor 1; + | intros; cases f; exists; [apply w] + [ assumption + | whd; intros; cases i; simplify; assumption] + | intros; split; intro; + [ cases f; cases x1; + (* senza questa change, universe inconsistency *) + change in ⊢ (? ? (λ_:%.?)) with (carr I); + exists [apply w1] exists [apply w] assumption; + | cases e; cases x; exists; [apply w1] + [ assumption + | (* senza questa change, universe inconsistency *) + whd; change in ⊢ (? ? (λ_:%.?)) with (carr I); + exists; [apply w] assumption]] + | intros; intros 2; cases (f (singleton ? 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 \sup -1‡#)); + assumption]] +qed.