]> matita.cs.unibo.it Git - helm.git/blob - matita/dama/integration_algebras.ma
Up to integration f-algebras.
[helm.git] / 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 abelian_group : Type \def
48  { carrier:> Type;
49    plus: carrier → carrier → carrier;
50    zero: carrier;
51    opp: carrier → carrier;
52    ag_abelian_group_properties: is_abelian_group ? plus zero opp
53  }.
54  
55 notation "0" with precedence 89
56 for @{ 'zero }.
57
58 interpretation "Ring zero" 'zero =
59  (cic:/matita/integration_algebras/zero.con _).
60
61 interpretation "Ring plus" 'plus a b =
62  (cic:/matita/integration_algebras/plus.con _ a b).
63
64 interpretation "Ring opp" 'uminus a =
65  (cic:/matita/integration_algebras/opp.con _ a).
66
67 definition minus ≝
68  λG:abelian_group.λa,b:G. a + -b.
69
70 interpretation "Ring minus" 'minus a b =
71  (cic:/matita/integration_algebras/minus.con _ a b).
72  
73 theorem plus_assoc: ∀G:abelian_group. associative ? (plus G).
74  intro;
75  apply (plus_assoc_ ? ? ? ? (ag_abelian_group_properties G)).
76 qed.
77
78 theorem plus_comm: ∀G:abelian_group. symmetric ? (plus G).
79  intro;
80  apply (plus_comm_ ? ? ? ? (ag_abelian_group_properties G)).
81 qed.
82
83 theorem zero_neutral: ∀G:abelian_group. left_neutral ? (plus G) 0.
84  intro;
85  apply (zero_neutral_ ? ? ? ? (ag_abelian_group_properties G)).
86 qed.
87
88 theorem opp_inverse: ∀G:abelian_group. left_inverse ? (plus G) 0 (opp G).
89  intro;
90  apply (opp_inverse_ ? ? ? ? (ag_abelian_group_properties G)).
91 qed.
92
93 lemma cancellationlaw: ∀G:abelian_group.∀x,y,z:G. x+y=x+z → y=z. 
94 intros;
95 generalize in match (eq_f ? ? (λa.-x +a) ? ? H);
96 intros; clear H;
97 rewrite < plus_assoc in H1;
98 rewrite < plus_assoc in H1;
99 rewrite > opp_inverse in H1;
100 rewrite > zero_neutral in H1;
101 rewrite > zero_neutral in H1;
102 assumption.
103 qed.
104
105 (****************************** rings *********************************)
106
107 record is_ring (G:abelian_group) (mult:G→G→G) (one:G) : Prop
108
109  {  (* multiplicative monoid properties *)
110     mult_assoc_: associative ? mult;
111     one_neutral_left_: left_neutral ? mult one;
112     one_neutral_right_: right_neutral ? mult one;
113     (* ring properties *)
114     mult_plus_distr_left_: distributive_left ? mult (plus G);
115     mult_plus_distr_right_: distributive_right ? mult (plus G);
116     not_eq_zero_one_: (0 ≠ one)
117  }.
118  
119 record ring : Type \def
120  { r_abelian_group:> abelian_group;
121    mult: r_abelian_group → r_abelian_group → r_abelian_group;
122    one: r_abelian_group;
123    r_ring_properties: is_ring r_abelian_group mult one
124  }.
125
126 theorem mult_assoc: ∀R:ring.associative ? (mult R).
127  intros;
128  apply (mult_assoc_ ? ? ? (r_ring_properties R)).
129 qed.
130
131 theorem one_neutral_left: ∀R:ring.left_neutral ? (mult R) (one R).
132  intros;
133  apply (one_neutral_left_ ? ? ? (r_ring_properties R)).
134 qed.
135
136 theorem one_neutral_right: ∀R:ring.right_neutral ? (mult R) (one R).
137  intros;
138  apply (one_neutral_right_ ? ? ? (r_ring_properties R)).
139 qed.
140
141 theorem mult_plus_distr_left: ∀R:ring.distributive_left ? (mult R) (plus R).
142  intros;
143  apply (mult_plus_distr_left_ ? ? ? (r_ring_properties R)).
144 qed.
145
146 theorem mult_plus_distr_right: ∀R:ring.distributive_right ? (mult R) (plus R).
147  intros;
148  apply (mult_plus_distr_right_ ? ? ? (r_ring_properties R)).
149 qed.
150
151 theorem not_eq_zero_one: ∀R:ring.0 ≠ one R.
152  intros;
153  apply (not_eq_zero_one_ ? ? ? (r_ring_properties R)).
154 qed.
155
156 interpretation "Ring mult" 'times a b =
157  (cic:/matita/integration_algebras/mult.con _ a b).
158
159 notation "1" with precedence 89
160 for @{ 'one }.
161
162 interpretation "Field one" 'one =
163  (cic:/matita/integration_algebras/one.con _).
164
165 lemma eq_mult_zero_x_zero: ∀R:ring.∀x:R.0*x=0.
166  intros;
167  generalize in match (zero_neutral R 0); intro;
168  generalize in match (eq_f ? ? (λy.y*x) ? ? H); intro; clear H;
169  rewrite > mult_plus_distr_right in H1;
170  generalize in match (eq_f ? ? (λy.-(0*x)+y) ? ? H1); intro; clear H1;
171  rewrite < plus_assoc in H;
172  rewrite > opp_inverse in H;
173  rewrite > zero_neutral in H;
174  assumption.
175 qed.
176
177 lemma eq_mult_x_zero_zero: ∀R:ring.∀x:R.x*0=0.
178 intros;
179 generalize in match (zero_neutral R 0);
180 intro;
181 generalize in match (eq_f ? ? (\lambda y.x*y) ? ? H); intro; clear H;
182 rewrite > mult_plus_distr_left in H1;
183 generalize in match (eq_f ? ? (\lambda y. (-(x*0)) +y) ? ? H1);intro;
184 clear H1;
185 rewrite < plus_assoc in H;
186 rewrite > opp_inverse in H;
187 rewrite > zero_neutral in H;
188 assumption.
189 qed.
190
191 record is_field (R:ring) (inv:∀x:R.x ≠ 0 → R) : Prop
192
193  {  (* multiplicative abelian properties *)
194     mult_comm_: symmetric ? (mult R);
195     (* multiplicative group properties *)
196     inv_inverse_: ∀x.∀p: x ≠ 0. mult ? (inv x p) x = 1
197  }.
198
199 lemma opp_opp: \forall R:ring. \forall x:R. (-(-x))=x.
200 intros; 
201 apply (cancellationlaw ? (-x) ? ?); 
202 rewrite > (opp_inverse R x); 
203 rewrite > plus_comm;
204 rewrite > opp_inverse;
205 reflexivity.
206 qed.
207
208
209 let rec sum (C:Type) (plus:C→C→C) (zero,one:C) (n:nat) on n  ≝
210  match n with
211   [ O ⇒ zero
212   | (S m) ⇒ plus one (sum C plus zero one m)
213   ].
214
215 record field : Type \def
216  { f_ring:> ring;
217    inv: ∀x:f_ring. x ≠ 0 → f_ring;
218    field_properties: is_field f_ring inv
219  }.
220  
221 theorem mult_comm: ∀F:field.symmetric ? (mult F).
222  intro;
223  apply (mult_comm_ ? ? (field_properties F)).
224 qed.
225
226 theorem inv_inverse: ∀F:field.∀x.∀p: x ≠ 0. mult ? (inv F x p) x = 1.
227  intro;
228  apply (inv_inverse_ ? ? (field_properties F)).
229 qed.
230
231 definition sum_field ≝
232  λF:field. sum ? (plus F) (zero F) (one F).
233  
234 record is_ordered_field_ch0 (F:field) (le:F→F→Prop) : Prop \def
235  { of_mult_compat: ∀a,b. le 0 a → le 0 b → le 0 (a*b);
236    of_plus_compat: ∀a,b,c. le a b → le (a+c) (b+c);
237    of_weak_tricotomy : ∀a,b. a≠b → le a b ∨ le b a;
238    (* 0 characteristics *)
239    of_char0: ∀n. n > O → sum ? (plus F) 0 1 n ≠ 0
240  }.
241  
242 record ordered_field_ch0 : Type \def
243  { of_field:> field;
244    of_le: of_field → of_field → Prop;
245    of_ordered_field_properties:> is_ordered_field_ch0 of_field of_le
246  }.
247
248 interpretation "Ordered field le" 'leq a b =
249  (cic:/matita/integration_algebras/of_le.con _ a b).
250  
251 definition lt \def λF:ordered_field_ch0.λa,b:F.a ≤ b ∧ a ≠ b.
252
253 interpretation "Ordered field lt" 'lt a b =
254  (cic:/matita/integration_algebras/lt.con _ a b).
255
256 lemma le_zero_x_to_le_opp_x_zero: ∀F:ordered_field_ch0.∀x:F. 0 ≤ x → -x ≤ 0.
257 intros;
258  generalize in match (of_plus_compat ? ? F ? ? (-x) H); intro;
259  rewrite > zero_neutral in H1;
260  rewrite > plus_comm in H1;
261  rewrite > opp_inverse in H1;
262  assumption.
263 qed.
264
265 lemma le_x_zero_to_le_zero_opp_x: ∀F:ordered_field_ch0.∀x:F. x ≤ 0 → 0 ≤ -x.
266  intros;
267  generalize in match (of_plus_compat ? ? F ? ? (-x) H); intro;
268  rewrite > zero_neutral in H1;
269  rewrite > plus_comm in H1;
270  rewrite > opp_inverse in H1;
271  assumption.
272 qed.
273
274 (*
275 lemma eq_opp_x_times_opp_one_x: ∀F:ordered_field_ch0.∀x:F.-x = -1*x.
276  intros;
277  
278 lemma not_eq_x_zero_to_lt_zero_mult_x_x:
279  ∀F:ordered_field_ch0.∀x:F. x ≠ 0 → 0 < x * x.
280  intros;
281  elim (of_weak_tricotomy ? ? ? ? ? ? ? ? F ? ? H);
282   [ generalize in match (le_x_zero_to_le_zero_opp_x F ? H1); intro;
283     generalize in match (of_mult_compat ? ? ? ? ? ? ? ?  F ? ? H2 H2); intro;
284 *)  
285
286 (* The ordering is not necessary. *)
287 axiom not_eq_sum_field_zero: ∀F:ordered_field_ch0.∀n. O<n → sum_field F n ≠ 0.
288
289 record is_vector_space (K: field) (G:abelian_group) (emult:K→G→G) : Prop
290
291  { vs_nilpotent: ∀v. emult 0 v = 0;
292    vs_neutral: ∀v. emult 1 v = v;
293    vs_distributive: ∀a,b,v. emult (a + b) v = (emult a v) + (emult b v);
294    vs_associative: ∀a,b,v. emult (a * b) v = emult a (emult b v)
295  }.
296
297 record vector_space (K:field): Type \def
298 { vs_abelian_group :> abelian_group;
299   emult: K → vs_abelian_group → vs_abelian_group;
300   vs_vector_space_properties :> is_vector_space K vs_abelian_group emult
301 }.
302
303 interpretation "Vector space external product" 'times a b =
304  (cic:/matita/integration_algebras/emult.con _ _ a b).
305
306 record is_lattice (C:Type) (join,meet:C→C→C) : Prop \def
307  { (* abelian semigroup properties *)
308    l_comm_j: symmetric ? join;
309    l_associative_j: associative ? join;
310    l_comm_m: symmetric ? meet;
311    l_associative_m: associative ? meet;
312    (* other properties *)
313    l_adsorb_j_m: ∀f,g. join f (meet f g) = f;
314    l_adsorb_m_j: ∀f,g. meet f (join f g) = f
315  }.
316
317 record lattice (C:Type) : Type \def
318  { join: C → C → C;
319    meet: C → C → C;
320    l_lattice_properties: is_lattice ? join meet
321  }.
322
323 definition le \def λC:Type.λL:lattice C.λf,g. meet ? L f g = f.
324
325 interpretation "Lattice le" 'leq a b =
326  (cic:/matita/integration_algebras/le.con _ _ a b).
327
328 definition carrier_of_lattice ≝
329  λC:Type.λL:lattice C.C.
330
331 record is_riesz_space (K:ordered_field_ch0) (V:vector_space K)
332  (L:lattice (Type_OF_vector_space ? V))
333 : Prop
334 \def
335  { rs_compat_le_plus: ∀f,g,h. le ? L f g → le ? L (f+h) (g+h);
336    rs_compat_le_times: ∀a:K.∀f. of_le ? 0 a → le ? L 0 f → le ? L 0 (a*f)
337  }.
338
339 record riesz_space (K:ordered_field_ch0) : Type \def
340  { rs_vector_space:> vector_space K;
341    rs_lattice:> lattice rs_vector_space;
342    rs_riesz_space_properties: is_riesz_space ? rs_vector_space rs_lattice
343  }.
344
345 definition absolute_value \def λK.λS:riesz_space K.λf.join ? S f (-f).   
346
347 record is_archimedean_riesz_space (K) (S:riesz_space K) : Prop
348 \def
349   { ars_archimedean: ∃u.∀n.∀a.∀p:n > O.
350      le ? S
351       (absolute_value ? S a)
352       (emult ? S
353         (inv ? (sum_field K n) (not_eq_sum_field_zero ? n p))
354         u) →
355      a = 0
356   }.
357
358 record archimedean_riesz_space (K:ordered_field_ch0) : Type \def
359  { ars_riesz_space:> riesz_space K;
360    ars_archimedean_property: is_archimedean_riesz_space ? ars_riesz_space
361  }.
362
363 record is_algebra (K: field) (V:vector_space K) (mult:V→V→V) (one:V) : Prop
364
365  { (* ring properties *)
366    a_ring: is_ring V mult one;
367    (* algebra properties *)
368    a_associative_left: ∀a,f,g. a * (mult f g) = mult (a * f) g;
369    a_associative_right: ∀a,f,g. a * (mult f g) = mult f (a * g)
370  }.
371
372 record algebra (K: field) (V:vector_space K) : Type \def
373  { a_mult: V → V → V;
374    a_one: V;
375    a_algebra_properties: is_algebra K V a_mult a_one
376  }.
377
378 interpretation "Algebra product" 'times a b =
379  (cic:/matita/integration_algebras/a_mult.con _ _ _ a b).
380
381 interpretation "Algebra one" 'one =
382  (cic:/matita/integration_algebras/a_one.con _ _ _).
383
384 definition ring_of_algebra ≝
385  λK.λV:vector_space K.λA:algebra ? V.
386   mk_ring V (a_mult ? ? A) (a_one ? ? A)
387    (a_ring ? ? ? ? (a_algebra_properties ? ? A)).
388
389 coercion cic:/matita/integration_algebras/ring_of_algebra.con.
390
391 record is_f_algebra (K) (S:archimedean_riesz_space K) (A:algebra ? S) : Prop
392 \def 
393 { compat_mult_le:
394    ∀f,g:S.
395     le ? S 0 f → le ? S 0 g → le ? S 0 (a_mult ? ? A f g);
396   compat_mult_meet:
397    ∀f,g,h:S.
398     meet ? S f g = 0 → meet ? S (a_mult ? ? A h f) g = 0
399 }.
400
401 record f_algebra (K:ordered_field_ch0) : Type \def 
402 { fa_archimedean_riesz_space:> archimedean_riesz_space K;
403   fa_algebra:> algebra ? fa_archimedean_riesz_space;
404   fa_f_algebra_properties: is_f_algebra ? fa_archimedean_riesz_space fa_algebra
405 }.
406
407 (* to be proved; see footnote 2 in the paper by Spitters *)
408 axiom symmetric_a_mult: ∀K.∀A:f_algebra K. symmetric ? (a_mult ? ? A).
409
410
411 definition tends_to : ∀F:ordered_field_ch0.∀f:nat→F.∀l:F.Prop.
412  alias symbol "leq" = "Ordered field le".
413  alias id "le" = "cic:/matita/nat/orders/le.ind#xpointer(1/1)".
414  apply
415   (λF:ordered_field_ch0.λf:nat → F.λl:F.
416     ∀n:nat.∃m:nat.∀j:nat. le m j →
417      l - (inv F (sum_field F (S n)) ?) ≤ f j ∧
418      f j ≤ l + (inv F (sum_field F (S n)) ?));
419  apply not_eq_sum_field_zero;
420  unfold;
421  auto new.
422 qed.
423
424 record is_integral (K) (A:f_algebra K) (I:Type_OF_f_algebra ? A→K) : Prop
425 \def
426  { i_positive: ∀f:Type_OF_f_algebra ? A. le ? (lattice_OF_f_algebra ? A) 0 f → of_le K 0 (I f);
427    i_linear1: ∀f,g:Type_OF_f_algebra ? A. I (f + g) = I f + I g;
428    i_linear2: ∀f:A.∀k:K. I (emult ? A k f) = k*(I f)
429  }.
430
431 (* Here we are avoiding a construction (the quotient space to define
432    f=g iff I(|f-g|)=0 *)
433 record is_integration_f_algebra (K) (A:f_algebra K) (I:Type_OF_f_algebra ? A→K) : Prop
434 \def
435  { ifa_integral: is_integral ? ? I;
436    ifa_limit1:
437     ∀f:A. tends_to ? (λn.I(meet ? A f ((sum_field K n)*(a_one ? ? A)))) (I f);
438    ifa_limit2:
439     ∀f:A.
440      tends_to ?
441       (λn.
442         I (meet ? A f
443          ((inv ? (sum_field K (S n))
444            (not_eq_sum_field_zero K (S n) (le_S_S O n (le_O_n n)))
445          ) * (a_one ? ? A)))) 0;
446    ifa_quotient_space1:
447     ∀f,g:A. f=g → I(absolute_value ? A (f - g)) = 0
448  }.