X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fcontribs%2Fdama%2Fdama%2Fattic%2Fordered_fields_ch0.ma;fp=matita%2Fcontribs%2Fdama%2Fdama%2Fattic%2Fordered_fields_ch0.ma;h=898148d6c4013b093e58ae42a11677538d620ead;hp=0000000000000000000000000000000000000000;hb=f61af501fb4608cc4fb062a0864c774e677f0d76;hpb=58ae1809c352e71e7b5530dc41e2bfc834e1aef1 diff --git a/matita/contribs/dama/dama/attic/ordered_fields_ch0.ma b/matita/contribs/dama/dama/attic/ordered_fields_ch0.ma new file mode 100644 index 000000000..898148d6c --- /dev/null +++ b/matita/contribs/dama/dama/attic/ordered_fields_ch0.ma @@ -0,0 +1,151 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + + + +include "attic/fields.ma". +include "ordered_group.ma". + +(*CSC: non capisco questi alias! Una volta non servivano*) +alias id "plus" = "cic:/matita/group/plus.con". +alias symbol "plus" = "Abelian group plus". + +record pre_ordered_field_ch0: Type ≝ + { of_field:> field; + of_ordered_abelian_group_: ordered_abelian_group; + of_cotransitively_ordered_set_: cotransitively_ordered_set; + of_with1_: + cos_ordered_set of_cotransitively_ordered_set_ = + og_ordered_set of_ordered_abelian_group_; + of_with2: + og_abelian_group of_ordered_abelian_group_ = r_abelian_group of_field + }. + +lemma of_ordered_abelian_group: pre_ordered_field_ch0 → ordered_abelian_group. + intro F; + apply mk_ordered_abelian_group; + [ apply mk_pre_ordered_abelian_group; + [ apply (r_abelian_group F) + | apply (og_ordered_set (of_ordered_abelian_group_ F)) + | apply (eq_f ? ? carrier); + apply (of_with2 F) + ] + | + apply + (eq_rect' ? ? + (λG:abelian_group.λH:og_abelian_group (of_ordered_abelian_group_ F)=G. + is_ordered_abelian_group + (mk_pre_ordered_abelian_group G (of_ordered_abelian_group_ F) + (eq_f abelian_group Type carrier (of_ordered_abelian_group_ F) G + H))) + ? ? (of_with2 F)); + simplify; + apply (og_ordered_abelian_group_properties (of_ordered_abelian_group_ F)) + ] +qed. + +coercion cic:/matita/attic/ordered_fields_ch0/of_ordered_abelian_group.con. + +(*CSC: I am not able to prove this since unfold is undone by coercion composition*) +axiom of_with1: + ∀G:pre_ordered_field_ch0. + cos_ordered_set (of_cotransitively_ordered_set_ G) = + og_ordered_set (of_ordered_abelian_group G). + +lemma of_cotransitively_ordered_set : pre_ordered_field_ch0 → cotransitively_ordered_set. + intro F; + apply mk_cotransitively_ordered_set; + [ apply (og_ordered_set F) + | apply + (eq_rect ? ? (λa:ordered_set.cotransitive (os_carrier a) (os_le a)) + ? ? (of_with1 F)); + apply cos_cotransitive + ] +qed. + +coercion cic:/matita/attic/ordered_fields_ch0/of_cotransitively_ordered_set.con. + +record is_ordered_field_ch0 (F:pre_ordered_field_ch0) : Type \def + { of_mult_compat: ∀a,b:F. 0≤a → 0≤b → 0≤a*b; + of_weak_tricotomy : ∀a,b:F. a≠b → a≤b ∨ b≤a; + (* 0 characteristics *) + of_char0: ∀n. n > O → sum ? (plus F) 0 1 n ≠ 0 + }. + +record ordered_field_ch0 : Type \def + { of_pre_ordered_field_ch0:> pre_ordered_field_ch0; + of_ordered_field_properties:> is_ordered_field_ch0 of_pre_ordered_field_ch0 + }. + +(* +lemma eq_opp_x_times_opp_one_x: ∀F:ordered_field_ch0.∀x:F.-x = -1*x. + intros; + +lemma not_eq_x_zero_to_lt_zero_mult_x_x: + ∀F:ordered_field_ch0.∀x:F. x ≠ 0 → 0 < x * x. + intros; + elim (of_weak_tricotomy ? ? ? ? ? ? ? ? F ? ? H); + [ generalize in match (le_x_zero_to_le_zero_opp_x F ? H1); intro; + 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. lt F 0 x → lt F 0 (inv ? x p). + +alias symbol "lt" = "natural 'less than'". + +(* The ordering is not necessary. *) +axiom not_eq_sum_field_zero: ∀F:ordered_field_ch0.∀n. O