X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fdama%2Fordered_groups.ma;h=fb4b29f0dc3dd59ff9ae31163ca4d2471a1ab6ad;hb=9ff984b29ac963eef2f79521ce9dd7cbb9ae2c59;hp=a9912d43f55ccb92bd22ecec4c9e3320e68146ef;hpb=06a089726af079d5b2fe42ba78632565dad0eb3e;p=helm.git diff --git a/matita/dama/ordered_groups.ma b/matita/dama/ordered_groups.ma index a9912d43f..fb4b29f0d 100644 --- a/matita/dama/ordered_groups.ma +++ b/matita/dama/ordered_groups.ma @@ -14,58 +14,91 @@ set "baseuri" "cic:/matita/ordered_groups/". +include "ordered_sets.ma". include "groups.ma". -include "ordered_sets2.ma". -record pre_ordered_abelian_group : Type ≝ - { og_abelian_group:> abelian_group; - og_ordered_set_: ordered_set; - og_with: os_carrier og_ordered_set_ = og_abelian_group - }. +record pre_ogroup : Type ≝ { + og_abelian_group_: abelian_group; + og_tordered_set:> tordered_set; + og_with: carr og_abelian_group_ = og_tordered_set +}. -lemma og_ordered_set: pre_ordered_abelian_group → ordered_set. - intro G; - apply mk_ordered_set; - [ apply (carrier (og_abelian_group G)) - | apply (eq_rect ? ? (λC:Type.C→C→Prop) ? ? (og_with G)); - apply os_le - | apply - (eq_rect' ? ? - (λa:Type.λH:os_carrier (og_ordered_set_ G) = a. - is_order_relation a - (eq_rect Type (og_ordered_set_ G) (λC:Type.C→C→Prop) - (os_le (og_ordered_set_ G)) a H)) - ? ? (og_with G)); - simplify; - apply (os_order_relation_properties (og_ordered_set_ G)) - ] +lemma og_abelian_group: pre_ogroup → abelian_group. +intro G; apply (mk_abelian_group G); [1,2,3: rewrite < (og_with G)] +[apply (plus (og_abelian_group_ G));|apply zero;|apply opp] +unfold apartness_OF_pre_ogroup; cases (og_with G); simplify; +[apply plus_assoc|apply plus_comm|apply zero_neutral|apply opp_inverse|apply plus_strong_ext] qed. -coercion cic:/matita/ordered_groups/og_ordered_set.con. +coercion cic:/matita/ordered_groups/og_abelian_group.con. -definition is_ordered_abelian_group ≝ - λG:pre_ordered_abelian_group. ∀f,g,h:G. f≤g → f+h≤g+h. -record ordered_abelian_group : Type ≝ - { og_pre_ordered_abelian_group:> pre_ordered_abelian_group; - og_ordered_abelian_group_properties: - is_ordered_abelian_group og_pre_ordered_abelian_group - }. +record ogroup : Type ≝ { + og_carr:> pre_ogroup; + fle_plusr: ∀f,g,h:og_carr. f≤g → f+h≤g+h +}. -lemma le_zero_x_to_le_opp_x_zero: ∀G:ordered_abelian_group.∀x:G.0 ≤ x → -x ≤ 0. - intros; - generalize in match (og_ordered_abelian_group_properties ? ? ? (-x) H); intro; - rewrite > zero_neutral in H1; - rewrite > plus_comm in H1; - rewrite > opp_inverse in H1; - assumption. +lemma plus_cancr_le: + ∀G:ogroup.∀x,y,z:G.x+z ≤ y + z → x ≤ y. +intros 5 (G x y z L); +apply (le_rewl ??? (0+x) (zero_neutral ??)); +apply (le_rewl ??? (x+0) (plus_comm ???)); +apply (le_rewl ??? (x+(-z+z))); [apply feq_plusl;apply opp_inverse;] +apply (le_rewl ??? (x+(z+ -z))); [apply feq_plusl;apply plus_comm;] +apply (le_rewl ??? (x+z+ -z)); [apply eq_symmetric; apply plus_assoc;] +apply (le_rewr ??? (0+y) (zero_neutral ??)); +apply (le_rewr ??? (y+0) (plus_comm ???)); +apply (le_rewr ??? (y+(-z+z))); [apply feq_plusl;apply opp_inverse;] +apply (le_rewr ??? (y+(z+ -z))); [apply feq_plusl;apply plus_comm;] +apply (le_rewr ??? (y+z+ -z)); [apply eq_symmetric; apply plus_assoc;] +apply (fle_plusr ??? (-z)); +assumption; qed. -lemma le_x_zero_to_le_zero_opp_x: ∀G:ordered_abelian_group.∀x:G. x ≤ 0 → 0 ≤ -x. - intros; - generalize in match (og_ordered_abelian_group_properties ? ? ? (-x) H); intro; - rewrite > zero_neutral in H1; - rewrite > plus_comm in H1; - rewrite > opp_inverse in H1; - assumption. +lemma fle_plusl: ∀G:ogroup. ∀f,g,h:G. f≤g → h+f≤h+g. +intros (G f g h); +apply (plus_cancr_le ??? (-h)); +apply (le_rewl ??? (f+h+ -h)); [apply feq_plusr;apply plus_comm;] +apply (le_rewl ??? (f+(h+ -h)) (plus_assoc ????)); +apply (le_rewl ??? (f+(-h+h))); [apply feq_plusl;apply plus_comm;] +apply (le_rewl ??? (f+0)); [apply feq_plusl; apply eq_symmetric; apply opp_inverse] +apply (le_rewl ??? (0+f) (plus_comm ???)); +apply (le_rewl ??? (f) (eq_symmetric ??? (zero_neutral ??))); +apply (le_rewr ??? (g+h+ -h)); [apply feq_plusr;apply plus_comm;] +apply (le_rewr ??? (g+(h+ -h)) (plus_assoc ????)); +apply (le_rewr ??? (g+(-h+h))); [apply feq_plusl;apply plus_comm;] +apply (le_rewr ??? (g+0)); [apply feq_plusl; apply eq_symmetric; apply opp_inverse] +apply (le_rewr ??? (0+g) (plus_comm ???)); +apply (le_rewr ??? (g) (eq_symmetric ??? (zero_neutral ??))); +assumption; +qed. + +lemma plus_cancl_le: + ∀G:ogroup.∀x,y,z:G.z+x ≤ z+y → x ≤ y. +intros 5 (G x y z L); +apply (le_rewl ??? (0+x) (zero_neutral ??)); +apply (le_rewl ??? ((-z+z)+x)); [apply feq_plusr;apply opp_inverse;] +apply (le_rewl ??? (-z+(z+x)) (plus_assoc ????)); +apply (le_rewr ??? (0+y) (zero_neutral ??)); +apply (le_rewr ??? ((-z+z)+y)); [apply feq_plusr;apply opp_inverse;] +apply (le_rewr ??? (-z+(z+y)) (plus_assoc ????)); +apply (fle_plusl ??? (-z)); +assumption; +qed. + + +lemma le_zero_x_to_le_opp_x_zero: + ∀G:ogroup.∀x:G.0 ≤ x → -x ≤ 0. +intros (G x Px); apply (plus_cancr_le ??? x); +apply (le_rewl ??? 0 (eq_symmetric ??? (opp_inverse ??))); +apply (le_rewr ??? x (eq_symmetric ??? (zero_neutral ??))); +assumption; +qed. + +lemma le_x_zero_to_le_zero_opp_x: + ∀G:ogroup.∀x:G. x ≤ 0 → 0 ≤ -x. +intros (G x Lx0); apply (plus_cancr_le ??? x); +apply (le_rewr ??? 0 (eq_symmetric ??? (opp_inverse ??))); +apply (le_rewl ??? x (eq_symmetric ??? (zero_neutral ??))); +assumption; qed.