X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Fsubsets.ma;h=3c855236be7cbb1c1a6473b995e8cf16828c9bd5;hb=6ee4fa0ba5f4b6601b62afd482d4f30bd2de2f91;hp=4bbe7fdbebd35ce17de26a3bbc5ef5f672503477;hpb=bd3b9cc44e65e316159ab56d3099949224413d66;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 4bbe7fdbe..3c855236b 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/subsets.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/subsets.ma @@ -16,7 +16,7 @@ include "categories.ma". record powerset_carrier (A: objs1 SET) : Type1 ≝ { mem_operator: unary_morphism1 A CPROP }. -definition subseteq_operator: ∀A: SET. powerset_carrier A → powerset_carrier A → CProp1 ≝ +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). @@ -41,7 +41,7 @@ qed. interpretation "powerset" 'powerset A = (powerset_setoid1 A). interpretation "subset construction" 'subset \eta.x = - (mk_powerset_carrier _ (mk_unary_morphism1 _ CPROP x _)). + (mk_powerset_carrier ? (mk_unary_morphism1 ? CPROP x ?)). definition mem: ∀A. binary_morphism1 A (Ω \sup A) CPROP. intros; @@ -56,7 +56,7 @@ definition mem: ∀A. binary_morphism1 A (Ω \sup A) CPROP. | apply s1; assumption]] qed. -interpretation "mem" 'mem a S = (fun21 ___ (mem _) a S). +interpretation "mem" 'mem a S = (fun21 ??? (mem ?) a S). definition subseteq: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP. intros; @@ -71,7 +71,7 @@ definition subseteq: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP. apply (transitive_subseteq_operator ???? s s4) ]] qed. -interpretation "subseteq" 'subseteq U V = (fun21 ___ (subseteq _) U V). +interpretation "subseteq" 'subseteq U V = (fun21 ??? (subseteq ?) U V). @@ -90,13 +90,13 @@ definition overlaps: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) CPROP. [ apply (λA:objs1 SET.λU,V:Ω \sup A.(exT2 ? (λx:A.?(*x*) ∈ U) (λx:A.?(*x*) ∈ V) : CProp0)) | intros; constructor 1; intro; cases e2; exists; [1,4: apply w] - [ apply (. #‡e); assumption - | apply (. #‡e1); assumption - | apply (. #‡(e \sup -1)); assumption; - | apply (. #‡(e1 \sup -1)); assumption]] + [ apply (. #‡e^-1); assumption + | apply (. #‡e1^-1); assumption + | apply (. #‡e); assumption; + | apply (. #‡e1); assumption]] qed. -interpretation "overlaps" 'overlaps U V = (fun21 ___ (overlaps _) U V). +interpretation "overlaps" 'overlaps U V = (fun21 ??? (overlaps ?) U V). definition intersects: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) (Ω \sup A). @@ -106,11 +106,11 @@ definition intersects: intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1; | intros; split; intros 2; simplify in f ⊢ %; - [ apply (. (#‡e)‡(#‡e1)); assumption - | apply (. (#‡(e \sup -1))‡(#‡(e1 \sup -1))); assumption]] + [ apply (. (#‡e^-1)‡(#‡e1^-1)); assumption + | apply (. (#‡e)‡(#‡e1)); assumption]] qed. -interpretation "intersects" 'intersects U V = (fun21 ___ (intersects _) U V). +interpretation "intersects" 'intersects U V = (fun21 ??? (intersects ?) U V). definition union: ∀A. binary_morphism1 (Ω \sup A) (Ω \sup A) (Ω \sup A). @@ -120,16 +120,16 @@ definition union: intros; simplify; apply (.= (e‡#)‡(e‡#)); apply refl1 | intros; split; intros 2; simplify in f ⊢ %; - [ apply (. (#‡e)‡(#‡e1)); assumption - | apply (. (#‡(e \sup -1))‡(#‡(e1 \sup -1))); assumption]] + [ apply (. (#‡e^-1)‡(#‡e1^-1)); assumption + | apply (. (#‡e)‡(#‡e1)); assumption]] qed. -interpretation "union" 'union U V = (fun21 ___ (union _) U V). +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; + [ apply (λa:A.{b | a =_0 b}); unfold setoid1_of_setoid; simplify; intros; simplify; split; intro; apply (.= e1); @@ -139,66 +139,29 @@ 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). +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 H; + apply ({x | ∀i:I. x ∈ c i}); + simplify; intros; split; intros; [ apply (. (e^-1‡#)); | apply (. e‡#); ] + apply f; | intros; split; intros 2; simplify in f ⊢ %; intro; - [ apply (. (#‡(e i))); apply f; - | apply (. (#‡(e i)\sup -1)); apply f]] + [ apply (. (#‡(e i)^-1)); apply f; + | apply (. (#‡e i)); apply f]] qed. -axiom (*definition*) big_union: +definition big_union: ∀A:SET.∀I:SET.unary_morphism2 (setoid1_of_setoid I ⇒ Ω \sup A) (setoid2_of_setoid1 (Ω \sup A)). -(* intros; constructor 1; + intros; constructor 1; [ intro; whd; whd in A; whd in I; - apply ({x | ∃i:I. x ∈ t i}); - simplify; intros; split; intros; cases H; clear H; exists; [1,3:apply w] - [ apply (. (e‡#)); | apply (. (e \sup -1‡#)); ] + apply ({x | ∃i:I. x ∈ c i }); + simplify; intros; split; intros; cases e1; clear e1; exists; [1,3:apply w] + [ apply (. (e^-1‡#)); | apply (. (e‡#)); ] 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". - - -axiom daemon: False. -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; 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; exists [apply w1] exists [apply w] assumption; - | cases H; cases x; exists; [apply w1] [assumption | exists; [apply w] assumption]]*) - | intros; intros 2; cases (H (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]] -cases daemon; -qed. \ No newline at end of file + [ apply (. (#‡(e w)^-1)); apply x; + | apply (. (#‡e w)); apply x]] +qed.