]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/subsets.ma
...
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / subsets.ma
index 9847d4c55d07c2f3d5f2f83c9179446c9a2793d3..95d0284dc05ccc5397688f29c3cd26fefc60ae68 100644 (file)
 
 include "categories.ma".
 
-record powerset_carrier (A: objs1 SET) : Type1 ≝ { mem_operator: unary_morphism1 A CPROP }.
+record powerset_carrier (A: objs1 SET) : Type1 ≝ { mem_operator: A ⇒_1 CPROP }.
+interpretation "powerset low" 'powerset A = (powerset_carrier A).
+notation "hvbox(a break ∈. b)" non associative with precedence 45 for @{ 'mem_low $a $b }.
+interpretation "memlow" 'mem_low a S = (fun11 ?? (mem_operator ? S) a).
 
-definition subseteq_operator: ∀A: SET. powerset_carrier A → powerset_carrier A → CProp1 ≝
- λA:objs1 SET.λU,V.∀a:A. mem_operator ? U a → mem_operator ? V a.
+definition subseteq_operator: ∀A: objs1 SET. Ω^A → Ω^A → CProp0 ≝
+ λA:objs1 SET.λU,V.∀a:A. a ∈. U → a ∈. V.
 
 theorem transitive_subseteq_operator: ∀A. transitive2 ? (subseteq_operator A).
  intros 6; intros 2;
@@ -41,9 +44,9 @@ 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.
+definition mem: ∀A. A × Ω^A ⇒_1 CPROP.
  intros;
  constructor 1;
   [ apply (λx,S. mem_operator ? S x)
@@ -56,9 +59,9 @@ 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.
+definition subseteq: ∀A. Ω^A × Ω^A ⇒_1 CPROP.
  intros;
  constructor 1;
   [ apply (λU,V. subseteq_operator ? U V)
@@ -71,65 +74,60 @@ 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).
 
-
-
-theorem subseteq_refl: ∀A.∀S:Ω \sup A.S ⊆ S.
+theorem subseteq_refl: ∀A.∀S:Ω^A.S ⊆ S.
  intros 4; assumption.
 qed.
 
-theorem subseteq_trans: ∀A.∀S1,S2,S3: Ω \sup A. S1 ⊆ S2 → S2 ⊆ S3 → S1 ⊆ S3.
+theorem subseteq_trans: ∀A.∀S1,S2,S3: Ω^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.
+definition overlaps: ∀A. Ω^A × Ω^A ⇒_1 CPROP.
  intros;
  constructor 1;
-  (* 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))
+  [ apply (λA:objs1 SET.λU,V:Ω^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).
+definition intersects: ∀A. Ω^A × Ω^A ⇒_1 Ω^A.
  intros;
  constructor 1;
   [ apply rule (λU,V. {x | x ∈ U ∧ x ∈ V });
     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).
+definition union: ∀A. Ω^A × Ω^A ⇒_1 Ω^A.
  intros;
  constructor 1;
   [ apply (λU,V. {x | x ∈ U ∨ x ∈ V });
     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).
+definition singleton: ∀A:setoid. A ⇒_1 Ω^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,77 +137,45 @@ 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).
+notation > "{ term 19 a : S }" with precedence 90 for @{fun11 ?? (singleton $S) $a}.
 
-definition big_intersects:
- ∀A:SET.∀I:SET.unary_morphism2 (setoid1_of_setoid I ⇒ Ω \sup A) (setoid2_of_setoid1 (Ω \sup A)).
+definition big_intersects: ∀A:SET.∀I:SET.(I ⇒_2 Ω^A) ⇒_2 Ω^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.
 
-(* senza questo exT "fresco", universe inconsistency *)
-inductive exT (A:Type) (P:A→CProp) : CProp ≝
-  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)).
+definition big_union: ∀A:SET.∀I:SET.(I ⇒_2 Ω^A) ⇒_2 Ω^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 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]]
+     [ apply (. (#‡(e w)^-1)); apply x;
+     | apply (. (#‡e w)); 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 H; 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 (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]]
-qed.
\ No newline at end of file
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋃) \below (\emsp) term 90 p)" 
+non associative with precedence 50 for @{ 'bigcup $p }.
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋃) \below (ident i ∈  I) break term 90 p)" 
+non associative with precedence 50 for @{ 'bigcup_mk (λ${ident i}:$I.$p) }.
+notation > "hovbox(⋃ f)" non associative with precedence 60 for @{ 'bigcup $f }.
+
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋂) \below (\emsp) term 90 p)" 
+non associative with precedence 50 for @{ 'bigcap $p }.
+notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (⋂) \below (ident i ∈  I) break term 90 p)" 
+non associative with precedence 50 for @{ 'bigcap_mk (λ${ident i}:$I.$p) }.
+notation > "hovbox(⋂ f)" non associative with precedence 60 for @{ 'bigcap $f }.
+
+interpretation "bigcup" 'bigcup f = (fun12 ?? (big_union ??) f).
+interpretation "bigcap" 'bigcap f = (fun12 ?? (big_intersects ??) f).
+interpretation "bigcup mk" 'bigcup_mk f = (fun12 ?? (big_union ??) (mk_unary_morphism2 ?? f ?)).
+interpretation "bigcap mk" 'bigcap_mk f = (fun12 ?? (big_intersects ??) (mk_unary_morphism2 ?? f ?)).