X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fdatatypes%2Fsubsets.ma;h=8c7963d67ea3a28b3073cadb844f27d1d25f713a;hb=4dc47c9675ffd5fa50296ffaa9b5997501518c98;hp=26b449d74fdcbe745ebbab0feb9cdbf66c078409;hpb=b195056adc77e652f59ec0b46afe277b150e12c8;p=helm.git diff --git a/helm/software/matita/library/datatypes/subsets.ma b/helm/software/matita/library/datatypes/subsets.ma index 26b449d74..8c7963d67 100644 --- a/helm/software/matita/library/datatypes/subsets.ma +++ b/helm/software/matita/library/datatypes/subsets.ma @@ -15,34 +15,128 @@ include "logic/cprop_connectives.ma". include "datatypes/categories.ma". -record powerset (A: Type) : Type ≝ { char: A → CProp }. - -interpretation "powerset" 'powerset A = (powerset A). - -interpretation "subset construction" 'subset \eta.x = (mk_powerset _ x). - -definition mem ≝ λA.λS:Ω \sup A.λx:A. match S with [mk_powerset c ⇒ c x]. - -interpretation "mem" 'mem a S = (mem _ S a). - -definition overlaps ≝ λA:Type.λU,V:Ω \sup A.exT2 ? (λa:A. a ∈ U) (λa.a ∈ V). - -interpretation "overlaps" 'overlaps U V = (overlaps _ U V). - -definition subseteq ≝ λA:Type.λU,V:Ω \sup A.∀a:A. a ∈ U → a ∈ V. - -interpretation "subseteq" 'subseteq U V = (subseteq _ U V). - -definition intersects ≝ λA:Type.λU,V:Ω \sup A.{a | a ∈ U ∧ a ∈ V}. - -interpretation "intersects" 'intersects U V = (intersects _ U V). - -definition union ≝ λA:Type.λU,V:Ω \sup A.{a | a ∈ U ∨ a ∈ V}. - -interpretation "union" 'union U V = (union _ U V). - -include "logic/equality.ma". - -definition singleton ≝ λA:Type.λa:A.{b | a=b}. - -interpretation "singleton" 'singl a = (singleton _ a). \ No newline at end of file +record powerset_carrier (A: setoid) : Type ≝ { mem_operator: A ⇒ CPROP }. + +definition subseteq_operator ≝ + λA:setoid.λU,V.∀a:A. mem_operator ? U a → mem_operator ? V a. + +theorem transitive_subseteq_operator: ∀A. transitive ? (subseteq_operator A). + intros 6; intros 2; + apply s1; apply s; + assumption. +qed. + +(* + +definition powerset_setoid: setoid → setoid1. + intros (T); + constructor 1; + [ apply (powerset_carrier T) + | constructor 1; + [ apply (λU,V. subseteq_operator ? U V ∧ subseteq_operator ? V U) + | simplify; intros; split; intros 2; assumption + | simplify; intros (x y H); cases H; split; assumption + | simplify; intros (x y z H H1); cases H; cases H1; split; + apply transitive_subseteq_operator; [1,4: apply y ] + assumption ]] +qed. + +interpretation "powerset" 'powerset A = (powerset_setoid A). + +interpretation "subset construction" 'subset \eta.x = + (mk_powerset_carrier _ (mk_unary_morphism _ CPROP x _)). + +definition mem: ∀A. binary_morphism1 A (Ω \sup A) CPROP. + intros; + constructor 1; + [ apply (λx,S. mem_operator ? S x) + | intros 5; + cases b; clear b; cases b'; clear b'; simplify; intros; + apply (trans1 ????? (prop_1 ?? u ?? H)); + cases H1; whd in s s1; + split; intro; + [ apply s; assumption + | apply s1; assumption]] +qed. + +interpretation "mem" 'mem a S = (fun1 ___ (mem _) a S). + +definition subseteq: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP. + intros; + constructor 1; + [ apply (λU,V. subseteq_operator ? U V) + | intros; + cases H; cases H1; + split; intros 1; + [ apply (transitive_subseteq_operator ????? s2); + apply (transitive_subseteq_operator ???? s1 s4) + | apply (transitive_subseteq_operator ????? s3); + apply (transitive_subseteq_operator ???? s s4) ]] +qed. + +interpretation "subseteq" 'subseteq U V = (fun1 ___ (subseteq _) U V). + +theorem subseteq_refl: ∀A.∀S:Ω \sup A.S ⊆ S. + intros 4; assumption. +qed. + +theorem subseteq_trans: ∀A.∀S1,S2,S3: Ω \sup A. S1 ⊆ S2 → S2 ⊆ S3 → S1 ⊆ S3. + intros; apply transitive_subseteq_operator; [apply S2] assumption. +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)) + | intros; + constructor 1; intro; cases H2; exists; [1,4: apply w] + [ apply (. #‡H); assumption + | apply (. #‡H1); assumption + | apply (. #‡(H \sup -1)); assumption; + | apply (. #‡(H1 \sup -1)); assumption]] +qed. + +interpretation "overlaps" 'overlaps U V = (fun1 ___ (overlaps _) U V). + +definition intersects: + ∀A. binary_morphism1 (powerset_setoid A) (powerset_setoid A) (powerset_setoid A). + intros; + constructor 1; + [ apply (λU,V. {x | x ∈ U ∧ x ∈ V }); + intros; simplify; apply (.= (H‡#)‡(H‡#)); apply refl1; + | intros; + split; intros 2; simplify in f ⊢ %; + [ apply (. (#‡H)‡(#‡H1)); assumption + | apply (. (#‡(H \sup -1))‡(#‡(H1 \sup -1))); assumption]] +qed. + +interpretation "intersects" 'intersects U V = (fun1 ___ (intersects _) U V). + +definition union: + ∀A. binary_morphism1 (powerset_setoid A) (powerset_setoid A) (powerset_setoid A). + intros; + constructor 1; + [ apply (λU,V. {x | x ∈ U ∨ x ∈ V }); + intros; simplify; apply (.= (H‡#)‡(H‡#)); apply refl1 + | intros; + split; intros 2; simplify in f ⊢ %; + [ apply (. (#‡H)‡(#‡H1)); assumption + | apply (. (#‡(H \sup -1))‡(#‡(H1 \sup -1))); assumption]] +qed. + +interpretation "union" 'union U V = (fun1 ___ (union _) U V). + +definition singleton: ∀A:setoid. unary_morphism A (Ω \sup A). + intros; constructor 1; + [ apply (λA:setoid.λa:A.{b | a=b}); + intros; simplify; + split; intro; + apply (.= H1); + [ apply H | apply (H \sup -1) ] + | intros; split; intros 2; simplify in f ⊢ %; apply trans; + [ apply a |4: apply a'] try assumption; apply sym; assumption] +qed. + +interpretation "singleton" 'singl a = (fun_1 __ (singleton _) a). + +*) \ No newline at end of file