X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fdatatypes%2Fsubsets.ma;h=7c9a13195f2ba6f07635f3067c5e4e942a07d095;hb=59945285cda6b39178eeffedb32a37d3141fe844;hp=5483dfa37f1fd648a38eca5727cd9cb9d05cc134;hpb=da03907a38982b8b45459213f2b9581accac5143;p=helm.git diff --git a/helm/software/matita/library/datatypes/subsets.ma b/helm/software/matita/library/datatypes/subsets.ma index 5483dfa37..7c9a13195 100644 --- a/helm/software/matita/library/datatypes/subsets.ma +++ b/helm/software/matita/library/datatypes/subsets.ma @@ -74,6 +74,14 @@ 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; @@ -116,12 +124,15 @@ qed. interpretation "union" 'union U V = (fun1 ___ (union _) U V). -definition singleton: ∀A:setoid. A → Ω \sup A. - apply (λA:setoid.λa:A.{b | a=b}); - intros; simplify; - split; intro; - apply (.= H1); - [ apply H | apply (H \sup -1) ] +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 = (singleton _ a). +interpretation "singleton" 'singl a = (fun_1 __ (singleton _) a).