]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/property_exhaustivity.ma
notation factored, coercion commant taking terms and not only URI
[helm.git] / helm / software / matita / contribs / dama / dama / property_exhaustivity.ma
index 58de86885f9acd06fdd88bde9fcb92435ca18a25..d04ec1acca886c457ae4370f3a5fcdc722120a9d 100644 (file)
@@ -23,39 +23,39 @@ definition exhaustive ≝
      (b is_decreasing → b is_lower_located → b is_cauchy).
 
 lemma segment_upperbound:
-  ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.u is_upper_bound ⌊n,fst (a n)⌋.
-intros 5; change with (fst (a n) ≤ u); cases (a n); cases H; assumption;
+  ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.u is_upper_bound ⌊n,\fst (a n)⌋.
+intros 5; change with (\fst (a n) ≤ u); cases (a n); cases H; assumption;
 qed.
 
 lemma segment_lowerbound:
-  ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.l is_lower_bound ⌊n,fst (a n)⌋.
-intros 5; change with (l ≤ fst (a n)); cases (a n); cases H; assumption;
+  ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.l is_lower_bound ⌊n,\fst (a n)⌋.
+intros 5; change with (l ≤ \fst (a n)); cases (a n); cases H; assumption;
 qed.
 
 lemma segment_preserves_uparrow:
   ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.∀x,h. 
-    ⌊n,fst (a n)⌋ ↑ x → a ↑ 〈x,h〉.
+    ⌊n,\fst (a n)⌋ ↑ x → a ↑ 〈x,h〉.
 intros; cases H (Ha Hx); split [apply Ha] cases Hx; 
 split; [apply H1] intros;
-cases (H2 (fst y)); [2: apply H3;] exists [apply w] assumption;
+cases (H2 (\fst y)); [2: apply H3;] exists [apply w] assumption;
 qed.
 
 lemma segment_preserves_downarrow:
   ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.∀x,h. 
-    ⌊n,fst (a n)⌋ ↓ x → a ↓ 〈x,h〉.
+    ⌊n,\fst (a n)⌋ ↓ x → a ↓ 〈x,h〉.
 intros; cases H (Ha Hx); split [apply Ha] cases Hx; 
 split; [apply H1] intros;
-cases (H2 (fst y));[2:apply H3]; exists [apply w] assumption;
+cases (H2 (\fst y));[2:apply H3]; exists [apply w] assumption;
 qed.
     
 (* Fact 2.18 *)
 lemma segment_cauchy:
   ∀C:ordered_uniform_space.∀l,u:C.∀a:sequence {[l,u]}.
-    a is_cauchy → ⌊n,fst (a n)⌋ is_cauchy.
+    a is_cauchy → ⌊n,\fst (a n)⌋ is_cauchy.
 intros 7; 
 alias symbol "pi1" (instance 3) = "pair pi1".
 alias symbol "pi2" = "pair pi2".
-apply (H (λx:{[l,u]} square.U 〈fst (fst x),fst (snd x)〉));
+apply (H (λx:{[l,u]} square.U 〈\fst (\fst x),\fst (\snd x)〉));
 (unfold segment_ordered_uniform_space; simplify);
 exists [apply U] split; [assumption;]
 intro; cases b; intros; simplify; split; intros; assumption;
@@ -65,7 +65,7 @@ qed.
 lemma restrict_uniform_convergence_uparrow:
   ∀C:ordered_uniform_space.property_sigma C →
     ∀l,u:C.exhaustive {[l,u]} →
-     ∀a:sequence {[l,u]}.∀x:C. ⌊n,fst (a n)⌋ ↑ x → 
+     ∀a:sequence {[l,u]}.∀x:C. ⌊n,\fst (a n)⌋ ↑ x → 
       x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges 〈x,h〉.
 intros; cases H2 (Ha Hx); clear H2; cases Hx; split;
 [1: split;
@@ -87,7 +87,7 @@ qed.
 lemma restrict_uniform_convergence_downarrow:
   ∀C:ordered_uniform_space.property_sigma C →
     ∀l,u:C.exhaustive {[l,u]} →
-     ∀a:sequence {[l,u]}.∀x:C. ⌊n,fst (a n)⌋ ↓ x → 
+     ∀a:sequence {[l,u]}.∀x:C. ⌊n,\fst (a n)⌋ ↓ x → 
       x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges 〈x,h〉.
 intros; cases H2 (Ha Hx); clear H2; cases Hx; split;
 [1: split;