]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/dama/integration_algebras.ma
c44d3e4822fa91fe15ed8fc9c708c9527946a2b2
[helm.git] / helm / software / matita / dama / integration_algebras.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/integration_algebras/".
16
17 include "higher_order_defs/functions.ma".
18 include "nat/nat.ma".
19 include "nat/orders.ma".
20
21 definition left_neutral \def λC,op.λe:C. ∀x:C. op e x = x.
22
23 definition right_neutral \def λC,op. λe:C. ∀x:C. op x e=x.
24
25 definition left_inverse \def λC,op.λe:C.λinv:C→C. ∀x:C. op (inv x) x = e.
26
27 definition right_inverse \def λC,op.λe:C.λ inv: C→ C. ∀x:C. op x (inv x)=e. 
28
29 definition distributive_left ≝
30  λA:Type.λf:A→A→A.λg:A→A→A.
31   ∀x,y,z. f x (g y z) = g (f x y) (f x z).
32
33 definition distributive_right ≝
34  λA:Type.λf:A→A→A.λg:A→A→A.
35   ∀x,y,z. f (g x y) z = g (f x z) (f y z).
36
37 record is_abelian_group (C:Type) (plus:C→C→C) (zero:C) (opp:C→C) : Prop \def
38  { (* abelian additive semigroup properties *)
39     plus_assoc: associative ? plus;
40     plus_comm: symmetric ? plus;
41     (* additive monoid properties *)
42     zero_neutral: left_neutral ? plus zero;
43     (* additive group properties *)
44     opp_inverse: left_inverse ? plus zero opp
45  }.
46
47 record is_ring (C:Type) (plus:C→C→C) (mult:C→C→C) (zero:C) (opp:C→C) : Prop
48
49  { (* abelian group properties *)
50     abelian_group:> is_abelian_group ? plus zero opp;
51     (* multiplicative semigroup properties *)
52     mult_assoc: associative ? mult;
53     (* ring properties *)
54     mult_plus_distr_left: distributive_left C mult plus;
55     mult_plus_distr_right: distributive_right C mult plus
56  }.
57  
58 record ring : Type \def
59  { r_carrier:> Type;
60    r_plus: r_carrier → r_carrier → r_carrier;
61    r_mult: r_carrier → r_carrier → r_carrier;
62    r_zero: r_carrier;
63    r_opp: r_carrier → r_carrier;
64    r_ring_properties:> is_ring ? r_plus r_mult r_zero r_opp
65  }.
66
67 notation "0" with precedence 89
68 for @{ 'zero }.
69
70 interpretation "Ring zero" 'zero =
71  (cic:/matita/integration_algebras/r_zero.con _).
72
73 interpretation "Ring plus" 'plus a b =
74  (cic:/matita/integration_algebras/r_plus.con _ a b).
75
76 interpretation "Ring mult" 'times a b =
77  (cic:/matita/integration_algebras/r_mult.con _ a b).
78
79 interpretation "Ring opp" 'uminus a =
80  (cic:/matita/integration_algebras/r_opp.con _ a).
81
82 lemma eq_mult_zero_x_zero: ∀R:ring.∀x:R.0*x=0.
83  intros;
84  generalize in match (zero_neutral ? ? ? ? R 0); intro;
85  generalize in match (eq_f ? ? (λy.y*x) ? ? H); intro; clear H;
86  rewrite > (mult_plus_distr_right ? ? ? ? ? R) in H1;
87  generalize in match (eq_f ? ? (λy.-(0*x)+y) ? ? H1); intro; clear H1;
88  rewrite < (plus_assoc ? ? ? ? R) in H;
89  rewrite > (opp_inverse ? ? ? ? R) in H;
90  rewrite > (zero_neutral ? ? ? ? R) in H;
91  assumption.
92 qed.
93
94 lemma eq_mult_x_zero_zero: ∀R:ring.∀x:R.x*0=0.
95 intros;
96 generalize in match (zero_neutral ? ? ? ? R 0);
97 intro;
98 generalize in match (eq_f ? ? (\lambda y.x*y) ? ? H); intro; clear H;
99 rewrite > (mult_plus_distr_left ? ? ? ? ? R) in H1;
100 generalize in match (eq_f ? ? (\lambda y. (-(x*0)) +y) ? ? H1);intro;
101 clear H1;
102 rewrite < (plus_assoc ? ? ? ? R) in H;
103 rewrite > (opp_inverse ? ? ? ? R) in H;
104 rewrite > (zero_neutral ? ? ? ? R) in H;
105 assumption.
106
107
108 record is_field (C:Type) (plus:C→C→C) (mult:C→C→C) (zero,one:C) (opp:C→C)
109  (inv:∀x:C.x ≠ zero →C)  : Prop
110
111  {  (* ring properties *)
112     ring_properties:> is_ring ? plus mult zero opp;
113     (* multiplicative abelian properties *)
114     mult_comm: symmetric ? mult;
115     (* multiplicative monoid properties *)
116     one_neutral: left_neutral ? mult one;
117     (* multiplicative group properties *)
118     inv_inverse: ∀x.∀p: x ≠ zero. mult (inv x p) x = one;
119     (* integral domain *)
120     not_eq_zero_one: zero ≠ one
121  }.
122
123 lemma cancellationlaw: \forall R:ring. \forall x,y,z:R. 
124 (x+y=x+z) \to (y=z). 
125 intros;
126 generalize in match (eq_f ? ? (\lambda a. (-x +a)) ? ? H);
127 intros; clear H;
128 rewrite < (plus_assoc ? ? ? ? R) in H1;
129 rewrite < (plus_assoc ? ? ? ? R) in H1;
130 rewrite > (opp_inverse ? ? ? ? R) in H1;
131 rewrite > (zero_neutral ? ? ? ? R) in H1;
132 rewrite > (zero_neutral ? ? ? ? R) in H1;
133 assumption.
134 qed.
135
136
137 lemma opp_opp: \forall R:ring. \forall x:R. (-(-x))=x.
138 intros; 
139 apply (cancellationlaw ? (-x) ? ?); 
140 rewrite  > (opp_inverse ? ? ? ? R (x)); 
141 rewrite > (plus_comm ? ? ? ? R);
142 rewrite > (opp_inverse ? ? ? ? R);
143 reflexivity.
144 qed.
145
146
147
148
149 let rec sum (C:Type) (plus:C→C→C) (zero,one:C) (n:nat) on n  ≝
150  match n with
151   [ O ⇒ zero
152   | (S m) ⇒ plus one (sum C plus zero one m)
153   ].
154
155 record field : Type \def
156  { f_ring:> ring;
157    one: f_ring;
158    inv: ∀x:f_ring. x ≠ 0 → f_ring;
159    field_properties:>
160     is_field ? (r_plus f_ring) (r_mult f_ring) (r_zero f_ring) one
161      (r_opp f_ring) inv
162  }.
163
164 definition sum_field ≝
165  λF:field. sum ? (r_plus F) (r_zero F) (one F).
166  
167 notation "1" with precedence 89
168 for @{ 'one }.
169
170 interpretation "Field one" 'one =
171  (cic:/matita/integration_algebras/one.con _).
172
173 record is_ordered_field_ch0 (C:Type) (plus,mult:C→C→C) (zero,one:C) (opp:C→C)
174  (inv:∀x:C.x ≠ zero → C) (le:C→C→Prop) : Prop \def
175  { (* field properties *)
176    of_is_field:> is_field C plus mult zero one opp inv;
177    of_mult_compat: ∀a,b. le zero a → le zero b → le zero (mult a b);
178    of_plus_compat: ∀a,b,c. le a b → le (plus a c) (plus b c);
179    of_weak_tricotomy : ∀a,b. a≠b → le a b ∨ le b a;
180    (* 0 characteristics *)
181    of_char0: ∀n. n > O → sum ? plus zero one n ≠ zero
182  }.
183  
184 record ordered_field_ch0 : Type \def
185  { of_field:> field;
186    of_le: of_field → of_field → Prop;
187    of_ordered_field_properties:>
188     is_ordered_field_ch0 ? (r_plus of_field) (r_mult of_field) (r_zero of_field)
189      (one of_field) (r_opp of_field) (inv of_field) of_le
190  }.
191
192 interpretation "Ordered field le" 'leq a b =
193  (cic:/matita/integration_algebras/of_le.con _ a b).
194  
195 definition lt \def λF:ordered_field_ch0.λa,b:F.a ≤ b ∧ a ≠ b.
196
197 interpretation "Ordered field lt" 'lt a b =
198  (cic:/matita/integration_algebras/lt.con _ a b).
199
200 (*incompleto
201 lemma le_zero_x_to_le_opp_x_zero: ∀F:ordered_field_ch0.∀x:F. 0 ≤ x → -x ≤ 0.
202 intros;
203  generalize in match (of_plus_compat ? ? ? ? ? ? ? ? F ? ? (-x) H); intro;
204  rewrite > (zero_neutral ? ? ? ? F) in H1;
205  rewrite > (plus_comm  ? ? ?  ? F) in H1;
206  rewrite > (opp_inverse ? ? ? ? F) in H1;
207  
208  assumption.
209 qed.*)
210
211 axiom le_x_zero_to_le_zero_opp_x: ∀F:ordered_field_ch0.∀x:F. x ≤ 0 → 0 ≤ -x.
212 (* intros;
213  generalize in match (of_plus_compat ? ? ? ? ? ? ? ? F ? ? (-x) H); intro;
214  rewrite > (zero_neutral ? ? ? ? F) in H1;
215  rewrite > (plus_comm ? ? ? ? F) in H1;
216  rewrite > (opp_inverse ? ? ? ? F) in H1;
217  assumption.
218 qed.*)
219
220 (*
221 lemma eq_opp_x_times_opp_one_x: ∀F:ordered_field_ch0.∀x:F.-x = -1*x.
222  intros;
223  
224 lemma not_eq_x_zero_to_lt_zero_mult_x_x:
225  ∀F:ordered_field_ch0.∀x:F. x ≠ 0 → 0 < x * x.
226  intros;
227  elim (of_weak_tricotomy ? ? ? ? ? ? ? ? F ? ? H);
228   [ generalize in match (le_x_zero_to_le_zero_opp_x F ? H1); intro;
229     generalize in match (of_mult_compat ? ? ? ? ? ? ? ?  F ? ? H2 H2); intro;
230 *)  
231
232 axiom not_eq_sum_field_zero: ∀F,n. n > O → sum_field F n ≠ 0.
233
234 record is_vector_space (K: field) (C:Type) (plus:C→C→C) (zero:C) (opp:C→C)
235  (emult:K→C→C) : Prop
236
237  { (* abelian group properties *)
238    vs_abelian_group: is_abelian_group ? plus zero opp;
239    (* other properties *)
240    vs_nilpotent: ∀v. emult 0 v = zero;
241    vs_neutral: ∀v. emult 1 v = v;
242    vs_distributive: ∀a,b,v. emult (a + b) v = plus (emult a v) (emult b v);
243    vs_associative: ∀a,b,v. emult (a * b) v = emult a (emult b v)
244  }.
245
246 record vector_space : Type \def
247 {vs_ :
248
249
250 }
251 record is_lattice (C:Type) (join,meet:C→C→C) : Prop \def
252  { (* abelian semigroup properties *)
253    l_comm_j: symmetric ? join;
254    l_associative_j: associative ? join;
255    l_comm_m: symmetric ? meet;
256    l_associative_m: associative ? meet;
257    (* other properties *)
258    l_adsorb_j_m: ∀f,g. join f (meet f g) = f;
259    l_adsorb_m_j: ∀f,g. meet f (join f g) = f
260  }.
261
262 (* This should be a let-in field of the riesz_space!!! *)
263 definition le_ \def λC.λmeet:C→C→C.λf,g. meet f g = f.
264
265 record is_riesz_space (K:ordered_field_ch0) (C:Type) (plus:C→C→C) (zero:C)
266  (opp:C→C) (emult:K→C→C) (join,meet:C→C→C) : Prop \def
267  { (* vector space properties *)
268    rs_vector_space: is_vector_space K C plus zero opp emult;
269    (* lattice properties *)
270    rs_lattice: is_lattice C join meet;
271    (* other properties *)
272    rs_compat_le_plus: ∀f,g,h. le_ ? meet f g → le_ ? meet (plus f h) (plus g h);
273    rs_compat_le_times: ∀a,f. 0≤a → le_ ? meet zero f → le_ ? meet zero (emult a f)  
274  }.
275  
276 definition absolute_value \def λC:Type.λopp.λjoin:C→C→C.λf.join f (opp f).   
277
278 record is_archimedean_riesz_space (K:ordered_field_ch0) (C:Type) (plus:C→C→C)
279  (zero:C) (opp:C→C) (emult:K→C→C) (join,meet:C→C→C)
280  :Prop \def
281   { ars_riesz_space: is_riesz_space ? ? plus zero opp emult join meet;
282     ars_archimedean: ∃u.∀n,a.∀p:n > O.
283      le_ C meet (absolute_value ? opp join a)
284       (emult (inv K (sum_field K n) (not_eq_sum_field_zero K n p)) u) →
285      a = zero
286   }.
287
288 record is_algebra (K: field) (C:Type) (plus:C→C→C) (zero:C) (opp:C→C)
289  (emult:K→C→C) (mult:C→C→C) : Prop
290
291  { (* vector space properties *)
292    a_vector_space_properties: is_vector_space ? ? plus zero opp emult;
293    (* ring properties *)
294    a_ring: is_ring ? plus mult zero opp;
295    (* algebra properties *)
296    a_associative_left: ∀a,f,g. emult a (mult f g) = mult (emult a f) g;
297    a_associative_right: ∀a,f,g. emult a (mult f g) = mult f (emult a g)
298  }.
299  
300  
301 record is_f_algebra (K: ordered_field_ch0) (C:Type) (plus: C\to C \to C) 
302 (zero:C) (opp: C \to C) (emult: Type_OF_ordered_field_ch0 K\to C\to C) (mult: C\to C\to C) 
303 (join,meet: C\to C\to C) : Prop
304 \def 
305 { archimedean_riesz_properties:> is_archimedean_riesz_space K C
306  plus zero opp emult join meet ;          
307 algebra_properties:> is_algebra ? ? plus zero opp emult mult;
308 compat_mult_le: \forall f,g: C. le_ ? meet zero f \to le_ ? meet zero g \to
309  le_ ? meet zero (mult f g);
310 compat_mult_meet: \forall f,g,h. meet f g = zero \to meet (mult h f) g = zero
311 }.
312
313 record f_algebra : Type \def 
314 {
315
316 }