]> matita.cs.unibo.it Git - helm.git/blob - matita/dama/ordered_groups.ma
snapshot
[helm.git] / matita / dama / ordered_groups.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/ordered_groups/".
16
17 include "ordered_sets.ma".
18 include "groups.ma".
19
20 record pre_ordered_abelian_group : Type ≝
21  { og_abelian_group_: abelian_group;
22    og_tordered_set:> tordered_set;
23    og_with: carr og_abelian_group_ = og_tordered_set
24  }.
25
26 lemma og_abelian_group: pre_ordered_abelian_group → abelian_group.
27 intro G; apply (mk_abelian_group G); [1,2,3: rewrite < (og_with G)]
28 [apply (plus (og_abelian_group_ G));|apply zero;|apply opp]
29 unfold apartness_OF_pre_ordered_abelian_group; cases (og_with G); simplify;
30 [apply plus_assoc|apply plus_comm|apply zero_neutral|apply opp_inverse|apply plus_strong_ext]
31 qed.
32
33 coercion cic:/matita/ordered_groups/og_abelian_group.con.
34
35 definition is_ordered_abelian_group ≝
36  λG:pre_ordered_abelian_group. ∀f,g,h:G. f≤g → f+h≤g+h.
37
38 record ordered_abelian_group : Type ≝
39  { og_pre_ordered_abelian_group:> pre_ordered_abelian_group;
40    og_ordered_abelian_group_properties:
41     is_ordered_abelian_group og_pre_ordered_abelian_group
42  }.
43
44 lemma le_rewl: ∀E:excedence.∀x,z,y:E. x ≈ y → x ≤ z → y ≤ z.
45 intros (E x z y); apply (le_transitive ???? ? H1); 
46 clear H1 z; unfold in H; unfold; intro H1; apply H; clear H; 
47 lapply ap_cotransitive;  
48 intros (G x z y); intro Eyz; 
49
50
51 lemma plus_cancr_le: 
52   ∀G:ordered_abelian_group.∀x,y,z:G.x+z ≤ y + z → x ≤ y.
53 intros 5 (G x y z L);
54
55  apply L; clear L; elim (exc_cotransitive ???z Exy);
56
57 lemma le_zero_x_to_le_opp_x_zero: 
58   ∀G:ordered_abelian_group.∀x:G.0 ≤ x → -x ≤ 0.
59 intros (G x Px);
60 generalize in match (og_ordered_abelian_group_properties ? ? ? (-x) Px); intro;
61 (* ma cazzo, qui bisogna rifare anche i gruppi con ≈ ? *)
62  rewrite > zero_neutral in H;
63  rewrite > plus_comm in H;
64  rewrite > opp_inverse in H;
65  assumption.
66 qed.
67
68 lemma le_x_zero_to_le_zero_opp_x: ∀G:ordered_abelian_group.∀x:G. x ≤ 0 → 0 ≤ -x.
69  intros;
70  generalize in match (og_ordered_abelian_group_properties ? ? ? (-x) H); intro;
71  rewrite > zero_neutral in H1;
72  rewrite > plus_comm in H1;
73  rewrite > opp_inverse in H1;
74  assumption.
75 qed.