]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/dama/reals.ma
Some clean-up here and there in dama (coercions removed, implicits added, etc.)
[helm.git] / matita / dama / reals.ma
index f587435a3b38ff4f406fc125dd492c22b61e8513..336bfe2e7c6fb4c084f40c352067fed21eff4940 100644 (file)
@@ -18,7 +18,7 @@ include "ordered_fields_ch0.ma".
 
 record is_real (F:ordered_field_ch0) : Type
 ≝
- { r_archimedean: ∀x:F. ∃n:nat. x ≤ (sum_field F n);
+ { r_archimedean: ∀x:F. ∃n:nat. x ≤ (sum_field ? n);
    r_complete: is_complete F  
  }.
 
@@ -27,16 +27,15 @@ record real: Type \def
    r_real_properties: is_real r_ordered_field_ch0
  }.
 
-(* serve l'esistenziale in CProp!
-definition lim: ∀R:real.∀f:nat→R.is_cauchy_seq R f → R.
+definition lim: ∀R:real.∀f:nat→R.is_cauchy_seq ? f → R.
  intros;
- elim H;
+ elim (r_complete ? (r_real_properties R) ? H);
+ exact a.
 qed.
-*)
 
 definition max_seq: ∀R:real.∀x,y:R. nat → R.
  intros (R x y);
- elim (to_cotransitive R (of_le R) R 0 (inv ? (sum_field R (S n)) ?) (x-y));
+ elim (to_cotransitive R (of_le R) R 0 (inv ? (sum_field ? (S n)) ?) (x-y));
   [ apply x
   | apply not_eq_sum_field_zero ;
     unfold;
@@ -48,7 +47,7 @@ qed.
 
 axiom daemon: False.
 
-theorem cauchy_max_seq: ∀R:real.∀x,y. is_cauchy_seq ? (max_seq R x y).
+theorem cauchy_max_seq: ∀R:real.∀x,y:R. is_cauchy_seq ? (max_seq ? x y).
  intros;
  unfold;
  intros;
@@ -70,12 +69,17 @@ theorem cauchy_max_seq: ∀R:real.∀x,y. is_cauchy_seq ? (max_seq R x y).
       rewrite > (plus_comm ? x (-x));
       rewrite > opp_inverse;
       split;
-       [ elim daemon (* da finire *)
+       [ apply (le_zero_x_to_le_opp_x_zero R ?); 
+         apply lt_zero_to_le_inv_zero
        | apply lt_zero_to_le_inv_zero
        ]
     | simplify;
       split;
-       [ elim daemon (* da finire *)
+       [  apply (to_transitive ? ? 
+       (of_total_order_relation ? ? R) ? 0 ?);
+         [ apply (le_zero_x_to_le_opp_x_zero R ?)
+         | assumption
+         ]        
        | assumption
        ]
     ]
@@ -87,7 +91,7 @@ theorem cauchy_max_seq: ∀R:real.∀x,y. is_cauchy_seq ? (max_seq R x y).
  (not_eq_sum_field_zero R (S m) (le_S_S O m (le_O_n m)))));
      [ simplify;
        split;
-       [ elim daemon (* da finire *)
+       [ elim daemon
        | generalize in match (le_zero_x_to_le_opp_x_zero R ? t1);
          intro;
          unfold minus in H1;
@@ -103,16 +107,18 @@ theorem cauchy_max_seq: ∀R:real.∀x,y. is_cauchy_seq ? (max_seq R x y).
        rewrite > (plus_comm ? y (-y));
        rewrite > opp_inverse;
        split;
-       [ elim daemon (* da finire *)
+       [ elim daemon
        | apply lt_zero_to_le_inv_zero
        ]
      ]
   ].
+  elim daemon.
 qed.
 
 definition max: ∀R:real.R → R → R.
  intros (R x y);
elim (r_complete ? (r_real_properties R) ? ?);
-  [|| apply (cauchy_max_seq R x y) ]
apply (lim R (max_seq R x y));
+ apply cauchy_max_seq.
 qed.
+
+definition abs \def λR:real.λx:R. max R x (-x).
\ No newline at end of file