]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/property_sigma.ma
after a PITA, lebergue is dualized!
[helm.git] / helm / software / matita / contribs / dama / dama / property_sigma.ma
index 241882bdd11fb5fff540b3832f5f7d09ad3daf03..9d9485f9c41c84eb6b595472ff6b5eaaa6f730a6 100644 (file)
@@ -20,7 +20,7 @@ alias num (instance 0) = "natural number".
 definition property_sigma ≝
   λC:ordered_uniform_space.
    ∀U.us_unifbase ? U → 
-     ∃V:sequence (C square → Prop).
+     ∃V:sequence (C squareB → Prop).
        (∀i.us_unifbase ? (V i)) ∧ 
        ∀a:sequence C.∀x:C.(a ↑ x ∨ a ↓ x) → 
          (∀n.∀i,j.n ≤ i → n ≤ j → V n 〈a i,a j〉) → U 〈a 0,x〉.
@@ -57,7 +57,7 @@ qed.
 lemma max_le_r: ∀n,m,z.max n m ≤ z → m ≤ z.
 intros; rewrite > sym_max in H; apply (max_le_l ??? H); 
 qed.
-    
+  
 (* Lemma 3.6 *)   
 lemma sigma_cauchy:
   ∀C:ordered_uniform_space.property_sigma C →
@@ -91,26 +91,29 @@ cut ((⌊x,(m x:nat)⌋ : sequence nat_ordered_set) is_increasing) as Hm1; [2:
 clearbody m; unfold spec in m Hm Hm1; clear spec;
 cut (⌊x,a (m x)⌋ ↑ l ∨ ⌊x,a (m x)⌋ ↓ l) as H10; [2: 
  cases H1;
- [ left; apply (selection_uparrow ?? Hm a l H4);
- | right; apply (selection_downarrow ?? Hm a l H4);]] 
+ [ left; apply (selection_uparrow ? Hm a l H4);
+ | right; apply (selection_downarrow ? Hm a l H4);]] 
 lapply (H9 ?? H10) as H11; [
   exists [apply (m 0:nat)] intros;
   cases H1;
     [cases H5; cases H7; apply (ous_convex ?? H3 ? H11 (H12 (m O)));
     |cases H5; cases H7; cases (us_phi4 ?? H3 〈(a (m O)),l〉); 
      lapply (H14 H11) as H11'; apply (ous_convex ?? H3 〈l,(a (m O))〉 H11' (H12 (m O)));]
-  simplify; repeat split; [1,6:intro X; cases (os_coreflexive ?? X)|*: try apply H12;]
-  [1:change with (a (m O) ≤ a i);
-     apply (trans_increasing ?? H6); intro; apply (le_to_not_lt ?? H4 H14);
-  |2:change with (a i ≤ a (m O));
-     apply (trans_decreasing ?? H6); intro; apply (le_to_not_lt ?? H4 H16);]]  
+  simplify; repeat split; 
+  [1,6: apply (le_reflexive l); 
+  |2,5: apply (H12 (\fst (m 0)));
+  |3,8: apply (H12 i);
+  |4:change with (a (m O) ≤ a i);
+     apply (trans_increasing a H6 (\fst (m 0)) i); intro; apply (le_to_not_lt ?? H4 H14);
+  |7:change with (a i ≤ a (m O));
+     apply (trans_decreasing ? H6 (\fst (m 0)) i); intro; apply (le_to_not_lt ?? H4 H16);]]  
 clear H10; intros (p q r); change with (w p 〈a (m q),a (m r)〉);
 generalize in match (refl_eq nat (m p)); 
 generalize in match (m p) in ⊢ (? ? ? % → %); intro X; cases X (w1 H15); clear X; 
 intros (H16); simplify in H16:(? ? ? %); destruct H16;
 apply H15; [3: apply le_n]
-[1: lapply (trans_increasing ?? Hm1 p q) as T; [apply not_lt_to_le; apply T;]
+[1: lapply (trans_increasing ? Hm1 p q) as T; [apply not_lt_to_le; apply T;]
     apply (le_to_not_lt p q H4);
-|2: lapply (trans_increasing ?? Hm1 p r) as T; [apply not_lt_to_le; apply T;]
+|2: lapply (trans_increasing ? Hm1 p r) as T; [apply not_lt_to_le; apply T;]
     apply (le_to_not_lt p r H5);]
 qed.