]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/dama/ordered_groups.ma
Some notes for Enrico.
[helm.git] / matita / dama / ordered_groups.ma
index 9e72cab5944b57d3ab0d45cd947aa936efb17e9c..cc7e739280cbbd58e277850e2c300e1fe350da88 100644 (file)
@@ -17,59 +17,86 @@ set "baseuri" "cic:/matita/ordered_groups/".
 include "ordered_sets.ma".
 include "groups.ma".
 
-record pre_ordered_abelian_group : Type ≝
{ og_abelian_group_: abelian_group;
-   og_tordered_set:> tordered_set;
-   og_with: carr og_abelian_group_ = og_tordered_set
- }.
+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_abelian_group: pre_ordered_abelian_group → abelian_group.
+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_ordered_abelian_group; cases (og_with G); simplify;
+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_abelian_group.con.
 
-definition is_ordered_abelian_group ≝
- λG:pre_ordered_abelian_group. ∀f,g,h:G. f≤g → f+h≤g+h.
+(* CSC: NO! Cosi' e' nel frammento negativo. Devi scriverlo con l'eccedenza!
+   Tutto il resto del file e' da rigirare nel frammento positivo.
+*)
+record ogroup : Type ≝ { 
+  og_carr:> pre_ogroup;
+  fle_plusr: ∀f,g,h:og_carr. 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
- }.
-
-lemma le_rewl: ∀E:excedence.∀x,z,y:E. x ≈ y → x ≤ z → y ≤ z.
-intros (E x z y); apply (le_transitive ???? ? H1); 
-clear H1 z; unfold in H; unfold; intro H1; apply H; clear H; 
-lapply ap_cotransitive;  
-intros (G x z y); intro Eyz; 
+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)) (opp_inverse ??));
+apply (le_rewl ??? (x+(z+ -z)) (plus_comm ??z));
+apply (le_rewl ??? (x+z+ -z) (plus_assoc ????));
+apply (le_rewr ??? (0+y) (zero_neutral ??));
+apply (le_rewr ??? (y+0) (plus_comm ???));
+apply (le_rewr ??? (y+(-z+z)) (opp_inverse ??));
+apply (le_rewr ??? (y+(z+ -z)) (plus_comm ??z));
+apply (le_rewr ??? (y+z+ -z) (plus_assoc ????));
+apply (fle_plusr ??? (-z) L);
+qed.
 
+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) (plus_comm ? f h));
+apply (le_rewl ??? (f+(h+ -h)) (plus_assoc ????));
+apply (le_rewl ??? (f+(-h+h)) (plus_comm ? h (-h)));
+apply (le_rewl ??? (f+0) (opp_inverse ??));
+apply (le_rewl ??? (0+f) (plus_comm ???));
+apply (le_rewl ??? (f) (zero_neutral ??));
+apply (le_rewr ??? (g+h+ -h) (plus_comm ? h ?));
+apply (le_rewr ??? (g+(h+ -h)) (plus_assoc ????));
+apply (le_rewr ??? (g+(-h+h)) (plus_comm ??h));
+apply (le_rewr ??? (g+0) (opp_inverse ??));
+apply (le_rewr ??? (0+g) (plus_comm ???));
+apply (le_rewr ??? (g) (zero_neutral ??) H);
+qed.
 
-lemma plus_cancr_le: 
-  ∀G:ordered_abelian_group.∀x,y,z:G.x+z ≤ y + z → x ≤ y.
+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) (opp_inverse ??));
+apply (le_rewl ??? (-z+(z+x)) (plus_assoc ????));
+apply (le_rewr ??? (0+y) (zero_neutral ??));
+apply (le_rewr ??? ((-z+z)+y) (opp_inverse ??));
+apply (le_rewr ??? (-z+(z+y)) (plus_assoc ????));
+apply (fle_plusl ??? (-z) L);
+qed.
 
- apply L; clear L; elim (exc_cotransitive ???z Exy);
 
 lemma le_zero_x_to_le_opp_x_zero: 
-  ∀G:ordered_abelian_group.∀x:G.0 ≤ x → -x ≤ 0.
-intros (G x Px);
-generalize in match (og_ordered_abelian_group_properties ? ? ? (-x) Px); intro;
-(* ma cazzo, qui bisogna rifare anche i gruppi con ≈ ? *)
- rewrite > zero_neutral in H;
- rewrite > plus_comm in H;
- rewrite > opp_inverse in H;
- assumption.
+  ∀G:ogroup.∀x:G.0 ≤ x → -x ≤ 0.
+intros (G x Px); apply (plus_cancr_le ??? x);
+apply (le_rewl ??? 0 (opp_inverse ??));
+apply (le_rewr ??? x (zero_neutral ??) Px);
 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 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 (opp_inverse ??));
+apply (le_rewl ??? x (zero_neutral ??));
+assumption; 
 qed.