]> matita.cs.unibo.it Git - helm.git/blob - matita/dama/ordered_groups.ma
d9be248d0c01f45094300432a1012c152d130746
[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_ogroup : 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_ogroup → 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_ogroup; 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 (* CSC: NO! Cosi' e' nel frammento negativo. Devi scriverlo con l'eccedenza!
36    Tutto il resto del file e' da rigirare nel frammento positivo.
37 *)
38 record ogroup : Type ≝ { 
39   og_carr:> pre_ogroup;
40   exc_canc_plusr: ∀f,g,h:og_carr. f+h ≰ g+h → f ≰ g
41 }.
42
43 lemma fexc_plusr: 
44   ∀G:ogroup.∀x,y,z:G. x ≰ y → x+z ≰ y + z.
45 intros 5 (G x y z L); apply (exc_canc_plusr ??? (-z));
46 apply (exc_rewl ??? (x + (z + -z)) (plus_assoc ????));
47 apply (exc_rewl ??? (x + (-z + z)) (plus_comm ??z));
48 apply (exc_rewl ??? (x+0) (opp_inverse ??));
49 apply (exc_rewl ??? (0+x) (plus_comm ???));
50 apply (exc_rewl ??? x (zero_neutral ??));
51 apply (exc_rewr ??? (y + (z + -z)) (plus_assoc ????));
52 apply (exc_rewr ??? (y + (-z + z)) (plus_comm ??z));
53 apply (exc_rewr ??? (y+0) (opp_inverse ??));
54 apply (exc_rewr ??? (0+y) (plus_comm ???));
55 apply (exc_rewr ??? y (zero_neutral ??) L);
56 qed.
57
58 coercion cic:/matita/ordered_groups/fexc_plusr.con nocomposites.
59
60 lemma exc_canc_plusl: ∀G:ogroup.∀f,g,h:G. h+f ≰ h+g → f ≰ g.
61 intros 5 (G x y z L); apply (exc_canc_plusr ??? z);
62 apply (exc_rewl ??? (z+x) (plus_comm ???));
63 apply (exc_rewr ??? (z+y) (plus_comm ???) L);
64 qed.
65
66 lemma fexc_plusl: 
67   ∀G:ogroup.∀x,y,z:G. x ≰ y → z+x ≰ z+y.
68 intros 5 (G x y z L); apply (exc_canc_plusl ??? (-z));
69 apply (exc_rewl ???? (plus_assoc ??z x));
70 apply (exc_rewr ???? (plus_assoc ??z y));
71 apply (exc_rewl ??? (0+x) (opp_inverse ??));
72 apply (exc_rewr ??? (0+y) (opp_inverse ??));
73 apply (exc_rewl ???? (zero_neutral ??));
74 apply (exc_rewr ???? (zero_neutral ??) L);
75 qed.
76
77 coercion cic:/matita/ordered_groups/fexc_plusl.con nocomposites.
78
79 lemma plus_cancr_le: 
80   ∀G:ogroup.∀x,y,z:G.x+z ≤ y + z → x ≤ y.
81 intros 5 (G x y z L);
82 apply (le_rewl ??? (0+x) (zero_neutral ??));
83 apply (le_rewl ??? (x+0) (plus_comm ???));
84 apply (le_rewl ??? (x+(-z+z)) (opp_inverse ??));
85 apply (le_rewl ??? (x+(z+ -z)) (plus_comm ??z));
86 apply (le_rewl ??? (x+z+ -z) (plus_assoc ????));
87 apply (le_rewr ??? (0+y) (zero_neutral ??));
88 apply (le_rewr ??? (y+0) (plus_comm ???));
89 apply (le_rewr ??? (y+(-z+z)) (opp_inverse ??));
90 apply (le_rewr ??? (y+(z+ -z)) (plus_comm ??z));
91 apply (le_rewr ??? (y+z+ -z) (plus_assoc ????));
92 intro H; apply L; clear L; apply (exc_canc_plusr ??? (-z) H);
93 qed.
94
95 lemma fle_plusl: ∀G:ogroup. ∀f,g,h:G. f≤g → h+f≤h+g.
96 intros (G f g h);
97 apply (plus_cancr_le ??? (-h));
98 apply (le_rewl ??? (f+h+ -h) (plus_comm ? f h));
99 apply (le_rewl ??? (f+(h+ -h)) (plus_assoc ????));
100 apply (le_rewl ??? (f+(-h+h)) (plus_comm ? h (-h)));
101 apply (le_rewl ??? (f+0) (opp_inverse ??));
102 apply (le_rewl ??? (0+f) (plus_comm ???));
103 apply (le_rewl ??? (f) (zero_neutral ??));
104 apply (le_rewr ??? (g+h+ -h) (plus_comm ? h ?));
105 apply (le_rewr ??? (g+(h+ -h)) (plus_assoc ????));
106 apply (le_rewr ??? (g+(-h+h)) (plus_comm ??h));
107 apply (le_rewr ??? (g+0) (opp_inverse ??));
108 apply (le_rewr ??? (0+g) (plus_comm ???));
109 apply (le_rewr ??? (g) (zero_neutral ??) H);
110 qed.
111
112 lemma plus_cancl_le: 
113   ∀G:ogroup.∀x,y,z:G.z+x ≤ z+y → x ≤ y.
114 intros 5 (G x y z L);
115 apply (le_rewl ??? (0+x) (zero_neutral ??));
116 apply (le_rewl ??? ((-z+z)+x) (opp_inverse ??));
117 apply (le_rewl ??? (-z+(z+x)) (plus_assoc ????));
118 apply (le_rewr ??? (0+y) (zero_neutral ??));
119 apply (le_rewr ??? ((-z+z)+y) (opp_inverse ??));
120 apply (le_rewr ??? (-z+(z+y)) (plus_assoc ????));
121 apply (fle_plusl ??? (-z) L);
122 qed.
123
124 lemma exc_opp_x_zero_to_exc_zero_x: 
125   ∀G:ogroup.∀x:G.-x ≰ 0 → 0 ≰ x.
126 intros (G x H); apply (exc_canc_plusr ??? (-x));
127 apply (exc_rewr ???? (plus_comm ???));
128 apply (exc_rewr ???? (opp_inverse ??));
129 apply (exc_rewl ???? (zero_neutral ??) H);
130 qed.
131   
132 lemma le_zero_x_to_le_opp_x_zero: 
133   ∀G:ogroup.∀x:G.0 ≤ x → -x ≤ 0.
134 intros (G x Px); apply (plus_cancr_le ??? x);
135 apply (le_rewl ??? 0 (opp_inverse ??));
136 apply (le_rewr ??? x (zero_neutral ??) Px);
137 qed.
138
139 lemma exc_zero_opp_x_to_exc_x_zero: 
140   ∀G:ogroup.∀x:G. 0 ≰ -x → x ≰ 0.
141 intros (G x H); apply (exc_canc_plusl ??? (-x));
142 apply (exc_rewr ???? (plus_comm ???));
143 apply (exc_rewl ???? (opp_inverse ??));
144 apply (exc_rewr ???? (zero_neutral ??) H);
145 qed.
146
147 lemma le_x_zero_to_le_zero_opp_x: 
148   ∀G:ogroup.∀x:G. x ≤ 0 → 0 ≤ -x.
149 intros (G x Lx0); apply (plus_cancr_le ??? x);
150 apply (le_rewr ??? 0 (opp_inverse ??));
151 apply (le_rewl ??? x (zero_neutral ??));
152 assumption; 
153 qed.