]> 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_le_eq: ∀E:excedence.∀x,y:E. x ≤ y → y ≤ x → x ≈ y.
45 intros 6 (E x y L1 L2 H); cases H; [apply (L1 H1)|apply (L2 H1)]
46 qed.
47
48 lemma unfold_apart: ∀E:excedence. ∀x,y:E. x ≰ y ∨ y ≰ x → x # y.
49 unfold apart_of_excedence; unfold apart; simplify; intros; assumption;
50 qed.
51
52 lemma le_rewl: ∀E:excedence.∀z,y,x:E. x ≈ y → x ≤ z → y ≤ z.
53 intros (E z y x Exy Lxz); apply (le_transitive ???? ? Lxz);
54 intro Xyz; apply Exy; apply unfold_apart; right; assumption;
55 qed.
56
57 lemma le_rewr: ∀E:excedence.∀z,y,x:E. x ≈ y → z ≤ x → z ≤ y.
58 intros (E z y x Exy Lxz); apply (le_transitive ???? Lxz);
59 intro Xyz; apply Exy; apply unfold_apart; left; assumption;
60 qed.
61
62 lemma plus_cancr_le: 
63   ∀G:ordered_abelian_group.∀x,y,z:G.x+z ≤ y + z → x ≤ y.
64 intros 5 (G x y z L);
65 apply (le_rewl ??? (0+x) (zero_neutral ??));
66 apply (le_rewl ??? (x+0) (plus_comm ???));
67 apply (le_rewl ??? (x+(-z+z))); [apply feq_plusl;apply opp_inverse;]
68 apply (le_rewl ??? (x+(z+ -z))); [apply feq_plusl;apply plus_comm;]
69 apply (le_rewl ??? (x+z+ -z)); [apply eq_symmetric; apply plus_assoc;]
70 apply (le_rewr ??? (0+y) (zero_neutral ??));
71 apply (le_rewr ??? (y+0) (plus_comm ???));
72 apply (le_rewr ??? (y+(-z+z))); [apply feq_plusl;apply opp_inverse;]
73 apply (le_rewr ??? (y+(z+ -z))); [apply feq_plusl;apply plus_comm;]
74 apply (le_rewr ??? (y+z+ -z)); [apply eq_symmetric; apply plus_assoc;]
75 apply (og_ordered_abelian_group_properties ??? (-z));
76 assumption;
77 qed.
78
79 lemma le_zero_x_to_le_opp_x_zero: 
80   ∀G:ordered_abelian_group.∀x:G.0 ≤ x → -x ≤ 0.
81 intros (G x Px); apply (plus_cancr_le ??? x);
82 apply (le_rewl ??? 0 (eq_symmetric ??? (opp_inverse ??)));
83 apply (le_rewr ??? x (eq_symmetric ??? (zero_neutral ??)));
84 assumption;
85 qed.
86
87 lemma le_x_zero_to_le_zero_opp_x: 
88   ∀G:ordered_abelian_group.∀x:G. x ≤ 0 → 0 ≤ -x.
89 intros (G x Lx0); apply (plus_cancr_le ??? x);
90 apply (le_rewr ??? 0 (eq_symmetric ??? (opp_inverse ??)));
91 apply (le_rewl ??? x (eq_symmetric ??? (zero_neutral ??)));
92 assumption; 
93 qed.