X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fordered_fields_ch0.ma;h=c7bf906f54afc3bd0f39670f024e6fcf5a0edbf4;hb=894c3e4038c93b484896e81132eae55046e47605;hp=d95108eb327faa8d6d382b3a4ed5be5971e48750;hpb=85fce74d6a0cd8e83e59895e81f1bc09d21ef4c2;p=helm.git diff --git a/helm/software/matita/dama/ordered_fields_ch0.ma b/helm/software/matita/dama/ordered_fields_ch0.ma index d95108eb3..c7bf906f5 100644 --- a/helm/software/matita/dama/ordered_fields_ch0.ma +++ b/helm/software/matita/dama/ordered_fields_ch0.ma @@ -16,8 +16,25 @@ set "baseuri" "cic:/matita/ordered_fields_ch0/". include "fields.ma". -record is_ordered_field_ch0 (F:field) (le:F→F→Prop) : Prop \def - { of_mult_compat: ∀a,b. le 0 a → le 0 b → le 0 (a*b); +record is_total_order_relation (C:Type) (le:C→C→Prop) : Type \def + { to_cotransitive: ∀x,y,z:C. le x y → le x z ∨ le z y; + to_antisimmetry: ∀x,y:C. le x y → le y x → x=y + }. + +theorem to_transitive: ∀C,le. is_total_order_relation C le → transitive ? le. + intros; + unfold transitive; + intros; + elim (to_cotransitive ? ? i ? ? z H); + [ assumption + | rewrite < (to_antisimmetry ? ? i ? ? H1 t); + assumption + ]. +qed. + +record is_ordered_field_ch0 (F:field) (le:F→F→Prop) : Type \def + { of_total_order_relation:> is_total_order_relation ? le; + of_mult_compat: ∀a,b. le 0 a → le 0 b → le 0 (a*b); of_plus_compat: ∀a,b,c. le a b → le (a+c) (b+c); of_weak_tricotomy : ∀a,b. a≠b → le a b ∨ le b a; (* 0 characteristics *) @@ -68,7 +85,52 @@ lemma not_eq_x_zero_to_lt_zero_mult_x_x: generalize in match (of_mult_compat ? ? ? ? ? ? ? ? F ? ? H2 H2); intro; *) +axiom lt_zero_to_lt_inv_zero: + ∀F:ordered_field_ch0.∀x:F.∀p:x≠0. 0 < x → 0 < inv ? x p. + (* The ordering is not necessary. *) axiom not_eq_sum_field_zero: ∀F:ordered_field_ch0.∀n. O