]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/dama/ordered_groups.ma
snapshot
[helm.git] / matita / dama / ordered_groups.ma
index 10e8f189ab0d4aeea549902fa4c2671372188f48..4d2e18e2838b081990d19dab8e8c573137036e93 100644 (file)
@@ -18,30 +18,19 @@ include "groups.ma".
 include "ordered_sets.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
+ { 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_ordered_abelian_group → 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;
+[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.
@@ -52,12 +41,14 @@ record ordered_abelian_group : Type ≝
     is_ordered_abelian_group og_pre_ordered_abelian_group
  }.
 
-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;
+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.
 qed.