X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Ftopology%2Figft.ma;h=7f4f5e3591724cfef8502d4659fbc1e3574a89c9;hb=e956eab1116ae48a298e3c6701f93178e53ab24f;hp=29e1f1af31995ddc7f1a17158b3ee6c07fcd8d8d;hpb=5bfea0766a39d05605b6e88508f5376a88351331;p=helm.git diff --git a/helm/software/matita/nlibrary/topology/igft.ma b/helm/software/matita/nlibrary/topology/igft.ma index 29e1f1af3..7f4f5e359 100644 --- a/helm/software/matita/nlibrary/topology/igft.ma +++ b/helm/software/matita/nlibrary/topology/igft.ma @@ -1,174 +1,301 @@ -include "logic/connectives.ma". +include "sets/sets.ma". -nrecord powerset (X : Type[0]) : Type[1] ≝ { char : X → CProp[0] }. - -interpretation "char" 'subset p = (mk_powerset ? p). - -interpretation "pwset" 'powerset a = (powerset a). - -interpretation "in" 'mem a X = (char ? X a). - -ndefinition subseteq ≝ λA.λU,V : Ω^A. - ∀x.x ∈ U → x ∈ V. - -interpretation "subseteq" 'subseteq u v = (subseteq ? u v). - -ndefinition overlaps ≝ λA.λU,V : Ω^A. - ∃x.x ∈ U ∧ x ∈ V. - -interpretation "overlaps" 'overlaps u v = (overlaps ? u v). -(* -ndefinition intersect ≝ λA.λu,v:Ω\sup A.{ y | y ∈ u ∧ y ∈ v }. - -interpretation "intersect" 'intersects u v = (intersect ? u v). -*) -nrecord axiom_set : Type[1] ≝ { - S:> Type[0]; +nrecord Ax : Type[1] ≝ { + S:> setoid; (* Type[0]; *) I: S → Type[0]; C: ∀a:S. I a → Ω ^ S }. -ndefinition cover_set ≝ λc:∀A:axiom_set.Ω^A → A → CProp[0].λA,C,U. +notation "𝐈 \sub( ❨a❩ )" non associative with precedence 70 for @{ 'I $a }. +notation "𝐂 \sub ( ❨a,\emsp i❩ )" non associative with precedence 70 for @{ 'C $a $i }. + +notation > "𝐈 term 90 a" non associative with precedence 70 for @{ 'I $a }. +notation > "𝐂 term 90 a term 90 i" non associative with precedence 70 for @{ 'C $a $i }. + +interpretation "I" 'I a = (I ? a). +interpretation "C" 'C a i = (C ? a i). + +ndefinition cover_set ≝ λc:∀A:Ax.Ω^A → A → CProp[0].λA,C,U. ∀y.y ∈ C → c A U y. +(* a \ltri b *) notation "hvbox(a break ◃ b)" non associative with precedence 45 -for @{ 'covers $a $b }. (* a \ltri b *) +for @{ 'covers $a $b }. interpretation "covers set temp" 'covers C U = (cover_set ?? C U). -ninductive cover (A : axiom_set) (U : Ω^A) : A → CProp[0] ≝ +ninductive cover (A : Ax) (U : Ω^A) : A → CProp[0] ≝ | creflexivity : ∀a:A. a ∈ U → cover ? U a -| cinfinity : ∀a:A. ∀i:I ? a. (C ? a i ◃ U) → cover ? U a. +| cinfinity : ∀a:A.∀i:𝐈 a. 𝐂 a i ◃ U → cover ? U a. napply cover; nqed. interpretation "covers" 'covers a U = (cover ? U a). interpretation "covers set" 'covers a U = (cover_set cover ? a U). -ndefinition fish_set ≝ λf:∀A:axiom_set.Ω^A → A → CProp[0]. +ndefinition fish_set ≝ λf:∀A:Ax.Ω^A → A → CProp[0]. λA,U,V. ∃a.a ∈ V ∧ f A U a. +(* a \ltimes b *) notation "hvbox(a break ⋉ b)" non associative with precedence 45 -for @{ 'fish $a $b }. (* a \ltimes b *) +for @{ 'fish $a $b }. interpretation "fish set temp" 'fish A U = (fish_set ?? U A). -ncoinductive fish (A : axiom_set) (F : Ω^A) : A → CProp[0] ≝ -| cfish : ∀a. a ∈ F → (∀i:I ? a.C A a i ⋉ F) → fish A F a. +ncoinductive fish (A : Ax) (F : Ω^A) : A → CProp[0] ≝ +| cfish : ∀a. a ∈ F → (∀i:𝐈 a .𝐂 a i ⋉ F) → fish A F a. napply fish; nqed. interpretation "fish set" 'fish A U = (fish_set fish ? U A). interpretation "fish" 'fish a U = (fish ? U a). -nlet corec fish_rec (A:axiom_set) (U: Ω^A) +nlet corec fish_rec (A:Ax) (U: Ω^A) (P: Ω^A) (H1: P ⊆ U) - (H2: ∀a:A. a ∈ P → ∀j: I ? a. C ? a j ≬ P): + (H2: ∀a:A. a ∈ P → ∀j: 𝐈 a. 𝐂 a j ≬ P): ∀a:A. ∀p: a ∈ P. a ⋉ U ≝ ?. #a; #p; napply cfish; ##[ napply H1; napply p; -##| #i; ncases (H2 a p i); #x; *; #xC; #xP; napply ex_intro; ##[napply x] - napply conj; ##[ napply xC ] napply (fish_rec ? U P); nassumption; +##| #i; ncases (H2 a p i); #x; *; #xC; #xP; @; ##[napply x] + @; ##[ napply xC ] napply (fish_rec ? U P); nassumption; +##] +nqed. + +notation "◃U" non associative with precedence 55 +for @{ 'coverage $U }. + +ndefinition coverage : ∀A:Ax.∀U:Ω^A.Ω^A ≝ λA,U.{ a | a ◃ U }. + +interpretation "coverage cover" 'coverage U = (coverage ? U). + +ndefinition cover_equation : ∀A:Ax.∀U,X:Ω^A.CProp[0] ≝ λA,U,X. + ∀a.a ∈ X ↔ (a ∈ U ∨ ∃i:𝐈 a.∀y.y ∈ 𝐂 a i → y ∈ X). + +ntheorem coverage_cover_equation : ∀A,U. cover_equation A U (◃U). +#A; #U; #a; @; #H; +##[ nelim H; #b; (* manca clear *) + ##[ #bU; @1; nassumption; + ##| #i; #CaiU; #IH; @2; @ i; #c; #cCbi; ncases (IH ? cCbi); + ##[ #E; @; napply E; + ##| #_; napply CaiU; nassumption; ##] ##] +##| ncases H; ##[ #E; @; nassumption] + *; #j; #Hj; @2 j; #w; #wC; napply Hj; nassumption; +##] +nqed. + +ntheorem coverage_min_cover_equation : + ∀A,U,W. cover_equation A U W → ◃U ⊆ W. +#A; #U; #W; #H; #a; #aU; nelim aU; #b; +##[ #bU; ncases (H b); #_; #H1; napply H1; @1; nassumption; +##| #i; #CbiU; #IH; ncases (H b); #_; #H1; napply H1; @2; @i; napply IH; +##] +nqed. + +notation "⋉F" non associative with precedence 55 +for @{ 'fished $F }. + +ndefinition fished : ∀A:Ax.∀F:Ω^A.Ω^A ≝ λA,F.{ a | a ⋉ F }. + +interpretation "fished fish" 'fished F = (fished ? F). + +ndefinition fish_equation : ∀A:Ax.∀F,X:Ω^A.CProp[0] ≝ λA,F,X. + ∀a. a ∈ X ↔ a ∈ F ∧ ∀i:𝐈 a.∃y.y ∈ 𝐂 a i ∧ y ∈ X. + +ntheorem fised_fish_equation : ∀A,F. fish_equation A F (⋉F). +#A; #F; #a; @; (* bug, fare case sotto diverso da farlo sopra *) #H; ncases H; +##[ #b; #bF; #H2; @ bF; #i; ncases (H2 i); #c; *; #cC; #cF; @c; @ cC; + napply cF; +##| #aF; #H1; @ aF; napply H1; +##] +nqed. + +ntheorem fised_max_fish_equation : ∀A,F,G. fish_equation A F G → G ⊆ ⋉F. +#A; #F; #G; #H; #a; #aG; napply (fish_rec … aG); +#b; ncases (H b); #H1; #_; #bG; ncases (H1 bG); #E1; #E2; nassumption; +nqed. + +nrecord nAx : Type[2] ≝ { + nS:> setoid; (*Type[0];*) + nI: nS → Type[0]; + nD: ∀a:nS. nI a → Type[0]; + nd: ∀a:nS. ∀i:nI a. nD a i → nS +}. + +(* +TYPE f A → B, g : B → A, f ∘ g = id, g ∘ g = id. + +a = b → I a = I b +*) + +notation "𝐃 \sub ( ❨a,\emsp i❩ )" non associative with precedence 70 for @{ 'D $a $i }. +notation "𝐝 \sub ( ❨a,\emsp i,\emsp j❩ )" non associative with precedence 70 for @{ 'd $a $i $j}. + +notation > "𝐃 term 90 a term 90 i" non associative with precedence 70 for @{ 'D $a $i }. +notation > "𝐝 term 90 a term 90 i term 90 j" non associative with precedence 70 for @{ 'd $a $i $j}. + +interpretation "D" 'D a i = (nD ? a i). +interpretation "d" 'd a i j = (nd ? a i j). +interpretation "new I" 'I a = (nI ? a). + +ndefinition image ≝ λA:nAx.λa:A.λi. { x | ∃j:𝐃 a i. x = 𝐝 a i j }. + +notation > "𝐈𝐦 [𝐝 term 90 a term 90 i]" non associative with precedence 70 for @{ 'Im $a $i }. +notation "𝐈𝐦 [𝐝 \sub ( ❨a,\emsp i❩ )]" non associative with precedence 70 for @{ 'Im $a $i }. + +interpretation "image" 'Im a i = (image ? a i). + +ndefinition Ax_of_nAx : nAx → Ax. +#A; @ A (nI ?); #a; #i; napply (𝐈𝐦 [𝐝 a i]); +nqed. + +ninductive sigma (A : Type[0]) (P : A → CProp[0]) : Type[0] ≝ + sig_intro : ∀x:A.P x → sigma A P. + +interpretation "sigma" 'sigma \eta.p = (sigma ? p). + +ndefinition nAx_of_Ax : Ax → nAx. +#A; @ A (I ?); +##[ #a; #i; napply (Σx:A.x ∈ 𝐂 a i); +##| #a; #i; *; #x; #_; napply x; +##] +nqed. + +ninductive Ord (A : nAx) : Type[0] ≝ + | oO : Ord A + | oS : Ord A → Ord A + | oL : ∀a:A.∀i.∀f:𝐃 a i → Ord A. Ord A. + +notation "Λ term 90 f" non associative with precedence 50 for @{ 'oL $f }. +notation "x+1" non associative with precedence 50 for @{'oS $x }. + +interpretation "ordinals Lambda" 'oL f = (oL ? ? ? f). +interpretation "ordinals Succ" 'oS x = (oS ? x). + +nlet rec famU (A : nAx) (U : Ω^A) (x : Ord A) on x : Ω^A ≝ + match x with + [ oO ⇒ U + | oS y ⇒ let Un ≝ famU A U y in Un ∪ { x | ∃i.𝐈𝐦[𝐝 x i] ⊆ Un} + | oL a i f ⇒ { x | ∃j.x ∈ famU A U (f j) } ]. + +notation < "term 90 U \sub (term 90 x)" non associative with precedence 50 for @{ 'famU $U $x }. +notation > "U ⎽ term 90 x" non associative with precedence 50 for @{ 'famU $U $x }. + +interpretation "famU" 'famU U x = (famU ? U x). + +ndefinition ord_coverage : ∀A:nAx.∀U:Ω^A.Ω^A ≝ λA,U.{ y | ∃x:Ord A. y ∈ famU ? U x }. + +ndefinition ord_cover_set ≝ λc:∀A:nAx.Ω^A → Ω^A.λA,C,U. + ∀y.y ∈ C → y ∈ c A U. + +interpretation "coverage new cover" 'coverage U = (ord_coverage ? U). +interpretation "new covers set" 'covers a U = (ord_cover_set ord_coverage ? a U). +interpretation "new covers" 'covers a U = (mem ? (ord_coverage ? U) a). + +ntheorem new_coverage_reflexive: + ∀A:nAx.∀U:Ω^A.∀a. a ∈ U → a ◃ U. +#A; #U; #a; #H; @ (oO A); napply H; +nqed. + +nlemma ord_subset: + ∀A:nAx.∀a:A.∀i,f,U.∀j:𝐃 a i.U⎽(f j) ⊆ U⎽(Λ f). +#A; #a; #i; #f; #U; #j; #b; #bUf; @ j; nassumption; +nqed. + +naxiom AC : ∀A,a,i,U.(∀j:𝐃 a i.∃x:Ord A.𝐝 a i j ∈ U⎽x) → (Σf.∀j:𝐃 a i.𝐝 a i j ∈ U⎽(f j)). + +naxiom setoidification : + ∀A:nAx.∀a,b:A.∀U.a=b → b ∈ U → a ∈ U. + +alias symbol "covers" = "new covers". +alias symbol "covers" = "new covers set". +alias symbol "covers" = "new covers". +ntheorem new_coverage_infinity: + ∀A:nAx.∀U:Ω^A.∀a:A. (∃i:𝐈 a. 𝐈𝐦[𝐝 a i] ◃ U) → a ◃ U. +#A; #U; #a; *; #i; #H; nnormalize in H; +ncut (∀y:𝐃 a i.∃x:Ord A.𝐝 a i y ∈ U⎽x); ##[ + #y; napply H; @ y; napply #; ##] #H'; +ncases (AC … H'); #f; #Hf; +ncut (∀j.𝐝 a i j ∈ U⎽(Λ f)); + ##[ #j; napply (ord_subset … f … (Hf j));##] #Hf'; +@ ((Λ f)+1); @2; nwhd; @i; #x; *; #d; #Hd; +napply (setoidification … Hd); napply Hf'; +nqed. + +(* move away *) +nlemma subseteq_union: ∀A.∀U,V,W:Ω^A.U ⊆ W → V ⊆ W → U ∪ V ⊆ W. +#A; #U; #V; #W; #H; #H1; #x; *; #Hx; ##[ napply H; ##| napply H1; ##] nassumption; +nqed. + +nlemma new_coverage_min : + ∀A:nAx.∀U:qpowerclass A.∀V.U ⊆ V → (∀a:A.∀i.𝐈𝐦[𝐝 a i] ⊆ V → a ∈ V) → ◃(pc ? U) ⊆ V. +#A; #U; #V; #HUV; #Im; #b; *; #o; ngeneralize in match b; nchange with ((pc ? U)⎽o ⊆ V); +nelim o; +##[ #b; #bU0; napply HUV; napply bU0; +##| #p; #IH; napply subseteq_union; ##[ nassumption; ##] + #x; *; #i; #H; napply (Im ? i); napply (subseteq_trans … IH); napply H; +##| #a; #i; #f; #IH; #x; *; #d; napply IH; ##] +nqed. + +nlet rec famF (A: nAx) (F : Ω^A) (x : Ord A) on x : Ω^A ≝ + match x with + [ oO ⇒ F + | oS o ⇒ let Fo ≝ famF A F o in Fo ∩ { x | ∀i:𝐈 x.∃j:𝐃 x i.𝐝 x i j ∈ Fo } + | oL a i f ⇒ { x | ∀j:𝐃 a i.x ∈ famF A F (f j) } + ]. + +interpretation "famF" 'famU U x = (famF ? U x). + +ndefinition ord_fished : ∀A:nAx.∀F:Ω^A.Ω^A ≝ λA,F.{ y | ∀x:Ord A. y ∈ F⎽x }. + +interpretation "fished new fish" 'fished U = (ord_fished ? U). +interpretation "new fish" 'fish a U = (mem ? (ord_fished ? U) a). + +ntheorem new_fish_antirefl: + ∀A:nAx.∀F:Ω^A.∀a. a ⋉ F → a ∈ F. +#A; #F; #a; #H; nlapply (H (oO ?)); #aFo; napply aFo; +nqed. + +nlemma co_ord_subset: + ∀A:nAx.∀F:Ω^A.∀a,i.∀f:𝐃 a i → Ord A.∀j. F⎽(Λ f) ⊆ F⎽(f j). +#A; #F; #a; #i; #f; #j; #x; #H; napply H; +nqed. + +naxiom AC_dual : + ∀A:nAx.∀a:A.∀i,F. (∀f:𝐃 a i → Ord A.∃x:𝐃 a i.𝐝 a i x ∈ F⎽(f x)) → ∃j:𝐃 a i.∀x:Ord A.𝐝 a i j ∈ F⎽x. + + +ntheorem new_fish_compatible: + ∀A:nAx.∀F:Ω^A.∀a. a ⋉ F → ∀i:𝐈 a.∃j:𝐃 a i.𝐝 a i j ⋉ F. +#A; #F; #a; #aF; #i; nnormalize; +napply AC_dual; #f; +nlapply (aF (Λf+1)); #aLf; +nchange in aLf with (a ∈ F⎽(Λ f) ∧ ∀i:𝐈 a.∃j:𝐃 a i.𝐝 a i j ∈ F⎽(Λ f)); +ncut (∃j:𝐃 a i.𝐝 a i j ∈ F⎽(f j));##[ + ncases aLf; #_; #H; nlapply (H i); *; #j; #Hj; @j; napply Hj;##] #aLf'; +napply aLf'; +nqed. + +(* move away *) +nlemma subseteq_intersection_l: ∀A.∀U,V,W:Ω^A.U ⊆ W ∨ V ⊆ W → U ∩ V ⊆ W. +#A; #U; #V; #W; *; #H; #x; *; #xU; #xV; napply H; nassumption; +nqed. + +nlemma subseteq_intersection_r: ∀A.∀U,V,W:Ω^A.W ⊆ U → W ⊆ V → W ⊆ U ∩ V. +#A; #U; #V; #W; #H1; #H2; #x; #Hx; @; ##[ napply H1; ##| napply H2; ##] nassumption; +nqed. + +ntheorem max_new_fished: + ∀A:nAx.∀G,F:Ω^A.G ⊆ F → (∀a.a ∈ G → ∀i.𝐈𝐦[𝐝 a i] ≬ G) → G ⊆ ⋉F. +#A; #G; #F; #GF; #H; #b; #HbG; #o; ngeneralize in match HbG; ngeneralize in match b; +nchange with (G ⊆ F⎽o); +nelim o; +##[ napply GF; +##| #p; #IH; napply (subseteq_intersection_r … IH); + #x; #Hx; #i; ncases (H … Hx i); #c; *; *; #d; #Ed; #cG; + @d; napply IH; napply (setoidification … Ed^-1); napply cG; +##| #a; #i; #f; #Hf; nchange with (G ⊆ { y | ∀x. y ∈ F⎽(f x) }); + #b; #Hb; #d; napply (Hf d); napply Hb; ##] nqed. -STOP - -theorem reflexivity: ∀A:axiom_set.∀a:A.∀V. a ∈ V → a ◃ V. - intros; - apply refl; - assumption. -qed. - -theorem transitivity: ∀A:axiom_set.∀a:A.∀U,V. a ◃ U → U ◃ V → a ◃ V. - intros; - apply (covers_elim ?? {a | a ◃ V} ??? H); simplify; intros; - [ cases H1 in H2; apply H2; - | apply infinity; - [ assumption - | constructor 1; - assumption]] -qed. - -theorem covers_elim2: - ∀A: axiom_set. ∀U:Ω \sup A.∀P: A → CProp. - (∀a:A. a ∈ U → P a) → - (∀a:A.∀V:Ω \sup A. a ◃ V → V ◃ U → (∀y. y ∈ V → P y) → P a) → - ∀a:A. a ◃ U → P a. - intros; - change with (a ∈ {a | P a}); - apply (covers_elim ?????? H2); - [ intros 2; simplify; apply H; assumption - | intros; - simplify in H4 ⊢ %; - apply H1; - [ apply (C ? a1 j); - | autobatch; - | assumption; - | assumption]] -qed. - -theorem coreflexivity: ∀A:axiom_set.∀a:A.∀V. a ⋉ V → a ∈ V. - intros; - cases H; - assumption. -qed. - -theorem cotransitivity: - ∀A:axiom_set.∀a:A.∀U,V. a ⋉ U → (∀b:A. b ⋉ U → b ∈ V) → a ⋉ V. - intros; - apply (fish_rec ?? {a|a ⋉ U} ??? H); simplify; intros; - [ apply H1; apply H2; - | cases H2 in j; clear H2; intro i; - cases (H4 i); clear H4; exists[apply a3] assumption] -qed. - -theorem compatibility: ∀A:axiom_set.∀a:A.∀U,V. a ⋉ V → a ◃ U → U ⋉ V. - intros; - generalize in match H; clear H; - apply (covers_elim ?? {a|a ⋉ V → U ⋉ V} ??? H1); - clear H1; simplify; intros; - [ exists [apply x] assumption - | cases H2 in j H H1; clear H2 a1; intros; - cases (H1 i); clear H1; apply (H3 a1); assumption] -qed. - -definition leq ≝ λA:axiom_set.λa,b:A. a ◃ {y|b=y}. - -interpretation "covered by one" 'leq a b = (leq ? a b). - -theorem leq_refl: ∀A:axiom_set.∀a:A. a ≤ a. - intros; - apply refl; - normalize; - reflexivity. -qed. - -theorem leq_trans: ∀A:axiom_set.∀a,b,c:A. a ≤ b → b ≤ c → a ≤ c. - intros; - unfold in H H1 ⊢ %; - apply (transitivity ???? H); - constructor 1; - intros; - normalize in H2; - rewrite < H2; - assumption. -qed. - -definition uparrow ≝ λA:axiom_set.λa:A.mk_powerset ? (λb:A. a ≤ b). - -interpretation "uparrow" 'uparrow a = (uparrow ? a). - -definition downarrow ≝ λA:axiom_set.λU:Ω \sup A.mk_powerset ? (λa:A. (↑a) ≬ U). - -interpretation "downarrow" 'downarrow a = (downarrow ? a). - -definition fintersects ≝ λA:axiom_set.λU,V:Ω \sup A.↓U ∩ ↓V. - -interpretation "fintersects" 'fintersects U V = (fintersects ? U V). - -record convergent_generated_topology : Type ≝ - { AA:> axiom_set; - convergence: ∀a:AA.∀U,V:Ω \sup AA. a ◃ U → a ◃ V → a ◃ (U ↓ V) - }.