X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fdama%2Freals.ma;h=87bbee5fd0d2e84de0a9e49d8ee066f4d826e8b8;hb=936f80cf031a7b034dd70fef49abb90e69f2e680;hp=9e458e29102df18e4d84b35767c7095fd110f85e;hpb=dd3157d36216486d914a97cfff7a9cd34f009ffe;p=helm.git diff --git a/matita/dama/reals.ma b/matita/dama/reals.ma index 9e458e291..87bbee5fd 100644 --- a/matita/dama/reals.ma +++ b/matita/dama/reals.ma @@ -35,7 +35,7 @@ qed. definition max_seq: ∀R:real.∀x,y:R. nat → R. intros (R x y); - elim (of_cotransitive R 0 (inv ? (sum_field ? (S n)) ?) (x-y)); + elim (cos_cotransitive R 0 (inv ? (sum_field ? (S n)) ?) (x-y)); [ apply x | apply not_eq_sum_field_zero ; unfold; @@ -147,7 +147,7 @@ lemma technical1: ∀R:real.lim R (to_zero R) (is_cauchy_seq_to_zero R) = 0. elim daemon. qed. -lemma abs_x_ge_O: \forall R: real. \forall x:R. (zero R) ≤ abs R x. +lemma abs_x_ge_O: ∀R:real.∀x:R. 0 ≤ abs ? x. intros; unfold abs; unfold max; @@ -157,7 +157,7 @@ lemma abs_x_ge_O: \forall R: real. \forall x:R. (zero R) ≤ abs R x. unfold to_zero; unfold max_seq; elim - (of_cotransitive R 0 + (cos_cotransitive R 0 (inv R (sum_field R (S n)) (not_eq_sum_field_zero R (S n) (le_S_S O n (le_O_n n)))) (x--x) (lt_zero_to_le_inv_zero R (S n)