]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/ordered_uniform.ma
Snapshot.
[helm.git] / helm / software / matita / contribs / dama / dama / ordered_uniform.ma
index bf0260a510ac775fede207a1760f26e371e36daf..51ecff3f3887c78ab3536c8cf1273c31422902d1 100644 (file)
@@ -30,7 +30,7 @@ unfold bishop_set_OF_ordered_uniform_space_;
 |5: cases (with_ X); simplify; apply (us_phi4 (ous_us_ X))]
 qed.
 
-coercion cic:/matita/dama/ordered_uniform/ous_unifspace.con.
+coercion ous_unifspace.
 
 record ordered_uniform_space : Type ≝ {
   ous_stuff :> ordered_uniform_space_;
@@ -51,7 +51,7 @@ lemma segment_square_of_ordered_set_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.
+coercion segment_square_of_ordered_set_square with 0 2.
 
 alias symbol "pi1" (instance 4) = "exT \fst".
 alias symbol "pi1" (instance 2) = "exT \fst".
@@ -59,7 +59,7 @@ 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)〉.
 
-coercion cic:/matita/dama/ordered_uniform/ordered_set_square_of_segment_square.con.
+coercion ordered_set_square_of_segment_square.
 
 lemma restriction_agreement : 
   ∀O:ordered_uniform_space.∀l,r:O.∀P:{[l,r]} square → Prop.∀OP:O square → Prop.Prop.
@@ -85,7 +85,7 @@ lemma invert_restriction_agreement:
   ∀O:ordered_uniform_space.∀l,r:O.
    ∀U:{[l,r]} square → Prop.∀u:O square → Prop.
     restriction_agreement ? l r U u →
-    restriction_agreement ? l r (inv U) (inv 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);]
@@ -96,18 +96,21 @@ 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)〉.
 
+(*
 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".
 lemma ss_of_bs: 
  ∀O:ordered_set.∀u,v:O.
   ∀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.〈≪\fst b,H1≫,≪\snd b,H2≫〉.
 
+(*
 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: 
   ∀O:ordered_uniform_space.∀u,v:O.ordered_uniform_space.