]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/ordered_uniform.ma
removed <_,_> notation second interpretation for dependent pair, since it used
[helm.git] / helm / software / matita / contribs / dama / dama / ordered_uniform.ma
index 8aacb5b59214bafbb89911b4ca3f9530823f06d8..bf0260a510ac775fede207a1760f26e371e36daf 100644 (file)
@@ -39,7 +39,7 @@ record ordered_uniform_space : Type ≝ {
 
 definition invert_os_relation ≝
   λC:ordered_set.λU:C square → Prop.
-    λx:C square. U 〈snd x,fst x〉.
+    λx:C square. U 〈\snd x,\fst x〉.
 
 interpretation "relation invertion" 'invert a = (invert_os_relation _ a).
 interpretation "relation invertion" 'invert_symbol = (invert_os_relation _).
@@ -47,17 +47,17 @@ interpretation "relation invertion" 'invert_appl a x = (invert_os_relation _ a x
 
 lemma segment_square_of_ordered_set_square: 
   ∀O:ordered_set.∀u,v:O.∀x:O square.
-   fst x ∈ [u,v] → snd x ∈ [u,v] → {[u,v]} square.
-intros; split; exists; [1: apply (fst x) |3: apply (snd x)] assumption;
+   \fst x ∈ [u,v] → \snd x ∈ [u,v] → {[u,v]} square.
+intros; split; exists; [1: apply (\fst x) |3: apply (\snd x)] assumption;
 qed.
 
 coercion cic:/matita/dama/ordered_uniform/segment_square_of_ordered_set_square.con 0 2.
 
-alias symbol "pi1" (instance 4) = "exT fst".
-alias symbol "pi1" (instance 2) = "exT fst".
+alias symbol "pi1" (instance 4) = "exT \fst".
+alias symbol "pi1" (instance 2) = "exT \fst".
 lemma ordered_set_square_of_segment_square : 
  ∀O:ordered_set.∀u,v:O.{[u,v]} square → O square ≝ 
-  λO:ordered_set.λu,v:O.λb:{[u,v]} square.〈fst(fst b),fst(snd b)〉.
+  λO:ordered_set.λu,v:O.λb:{[u,v]} square.〈\fst(\fst b),\fst(\snd b)〉.
 
 coercion cic:/matita/dama/ordered_uniform/ordered_set_square_of_segment_square.con.
 
@@ -87,28 +87,26 @@ lemma invert_restriction_agreement:
     restriction_agreement ? l r U u →
     restriction_agreement ? l r (inv U) (inv u).
 intros 9; split; intro;
-[1: apply (unrestrict ????? (segment_square_of_ordered_set_square ??? 〈snd b,fst b〉 H2 H1) H H3);
-|2: apply (restrict ????? (segment_square_of_ordered_set_square ??? 〈snd b,fst b〉 H2 H1) H H3);]
+[1: apply (unrestrict ????? (segment_square_of_ordered_set_square ??? 〈\snd b,\fst b〉 H2 H1) H H3);
+|2: apply (restrict ????? (segment_square_of_ordered_set_square ??? 〈\snd b,\fst b〉 H2 H1) H H3);]
 qed. 
     
 alias symbol "square" (instance 8) = "bishop set square".
 lemma bs_of_ss: 
  ∀O:ordered_set.∀u,v:O.{[u,v]} square → (bishop_set_of_ordered_set O) square ≝ 
-  λO:ordered_set.λu,v:O.λb:{[u,v]} square.〈fst(fst b),fst(snd b)〉.
+  λO:ordered_set.λu,v:O.λb:{[u,v]} square.〈\fst(\fst b),\fst(\snd b)〉.
 
-notation < "x \sub \neq" left associative with precedence 91 for @{'bsss $x}.
+notation < "x \sub \neq" with precedence 91 for @{'bsss $x}.
 interpretation "bs_of_ss" 'bsss x = (bs_of_ss _ _ _ x).
 
 alias symbol "square" (instance 7) = "ordered set square".
-alias symbol "pair" (instance 4) = "dependent pair".
-alias symbol "pair" (instance 2) = "dependent pair".
 lemma ss_of_bs: 
  ∀O:ordered_set.∀u,v:O.
-  ∀b:O square.fst b ∈ [u,v] → snd b ∈ [u,v] → {[u,v]} square ≝ 
+  ∀b:O square.\fst b ∈ [u,v] → \snd b ∈ [u,v] → {[u,v]} square ≝ 
  λO:ordered_set.λu,v:O.
-  Î»b:(O:bishop_set) square.λH1,H2.â\8c©â\8c©fst b,H1â\8cª,â\8c©snd b,H2â\8cª〉.
+  Î»b:(O:bishop_set) square.λH1,H2.â\8c©â\89ª\fst b,H1â\89«,â\89ª\snd b,H2â\89«〉.
 
-notation < "x \sub \nleq" left associative with precedence 91 for @{'ssbs $x}.
+notation < "x \sub \nleq" with precedence 91 for @{'ssbs $x}.
 interpretation "ss_of_bs" 'ssbs x = (ss_of_bs _ _ _ x _ _).
 
 lemma segment_ordered_uniform_space: 
@@ -140,7 +138,7 @@ intros (O l r); apply mk_ordered_uniform_space;
                 unfold segment_square_of_ordered_set_square;
                 cases b; intros; split; intro; assumption;
             |2: intros 2 (x Hx); apply (restrict ?????? HuU); apply Hwu; 
-                cases Hx (m Hm); exists[apply (fst m)] apply Hm;]
+                cases Hx (m Hm); exists[apply (\fst m)] apply Hm;]
         |4: intros (U HU x); cases HU (u Hu); cases Hu (Gu HuU); clear HU Hu;
             cases (us_phi4 ?? Gu x) (Hul Hur);
             split; intros; 
@@ -161,12 +159,12 @@ interpretation "Ordered uniform space segment" 'segment_set a b =
  (segment_ordered_uniform_space _ a b).
 
 (* Lemma 3.2 *)
-alias symbol "pi1" = "exT fst".
+alias symbol "pi1" = "exT \fst".
 lemma restric_uniform_convergence:
  ∀O:ordered_uniform_space.∀l,u:O.
   ∀x:{[l,u]}.
    ∀a:sequence {[l,u]}.
-     ⌊n,fst (a n)⌋ uniform_converges (fst x) → 
+     ⌊n,\fst (a n)⌋ uniform_converges (\fst x) → 
       a uniform_converges x.
 intros 8; cases H1; cases H2; clear H2 H1;
 cases (H ? H3) (m Hm); exists [apply m]; intros;