]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/subsets.ma
few more files, one diverges
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / subsets.ma
index 3aaf2d3d6b34b6da73f52a678e6cfb7d95c5f73b..3c855236be7cbb1c1a6473b995e8cf16828c9bd5 100644 (file)
@@ -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,29 +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 ({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.
 
 definition big_union:
  ∀A:SET.∀I:SET.unary_morphism2 (setoid1_of_setoid I ⇒ Ω \sup A) (setoid2_of_setoid1 (Ω \sup A)).
  intros; constructor 1;
   [ intro; whd; whd in A; whd in I;
-    apply ({x | ∃i:carr I. x ∈ t i });
+    apply ({x | ∃i:I. x ∈ c i });
     simplify; intros; split; intros; cases e1; clear e1; exists; [1,3:apply w]
-    [ apply (. (e‡#)); | apply (. (e \sup -1‡#)); ]
+    [ 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.